diff --git a/src/Init/Meta.lean b/src/Init/Meta.lean index c5d02cd7c280..a5bd7344e2b6 100644 --- a/src/Init/Meta.lean +++ b/src/Init/Meta.lean @@ -158,8 +158,6 @@ protected def reprPrec (n : Name) (prec : Nat) : Std.Format := instance : Repr Name where reprPrec := Name.reprPrec -deriving instance Repr for Syntax - def capitalize : Name → Name | .str p s => .str p s.capitalize | n => n @@ -257,6 +255,9 @@ instance monadNameGeneratorLift (m n : Type → Type) [MonadLift m n] [MonadName namespace Syntax +deriving instance Repr for Syntax.Preresolved +deriving instance Repr for Syntax + abbrev Term := TSyntax `term abbrev Command := TSyntax `command protected abbrev Level := TSyntax `level @@ -325,6 +326,9 @@ end TSyntax namespace Syntax +deriving instance BEq for Syntax.Preresolved + +/-- Compare syntax structures modulo source info. -/ partial def structEq : Syntax → Syntax → Bool | Syntax.missing, Syntax.missing => true | Syntax.node _ k args, Syntax.node _ k' args' => k == k' && args.isEqv args' structEq @@ -510,7 +514,7 @@ def mkIdentFromRef [Monad m] [MonadRef m] (val : Name) : m Ident := do def mkCIdentFrom (src : Syntax) (c : Name) : Ident := -- Remark: We use the reserved macro scope to make sure there are no accidental collision with our frontend let id := addMacroScope `_internal c reservedMacroScope - ⟨Syntax.ident (SourceInfo.fromRef src) (toString id).toSubstring id [(c, [])]⟩ + ⟨Syntax.ident (SourceInfo.fromRef src) (toString id).toSubstring id [.decl c []]⟩ def mkCIdentFromRef [Monad m] [MonadRef m] (c : Name) : m Syntax := do return mkCIdentFrom (← getRef) c diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index 4dd161120898..59654d8a2683 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -3399,6 +3399,17 @@ abbrev SyntaxNodeKind := Name /-! # Syntax AST -/ +/-- +Binding information resolved and stored at compile time of a syntax quotation. +Note: We do not statically know whether a syntax expects a namespace or term name, +so a `Syntax.ident` may contain both preresolution kinds. +-/ +inductive Syntax.Preresolved where + | /-- A potential namespace reference -/ + namespace (ns : Name) + | /-- A potential global constant or section variable reference, with additional field accesses -/ + decl (n : Name) (fields : List String) + /-- Syntax objects used by the parser, macro expander, delaborator, etc. -/ @@ -3435,9 +3446,9 @@ inductive Syntax where `rawIdent` parsers. * `rawVal` is the literal substring from the input file * `val` is the parsed identifier (with hygiene) - * `preresolved` is the list of possible constants this could refer to + * `preresolved` is the list of possible declarations this could refer to -/ - ident (info : SourceInfo) (rawVal : Substring) (val : Name) (preresolved : List (Prod Name (List String))) : Syntax + ident (info : SourceInfo) (rawVal : Substring) (val : Name) (preresolved : List Syntax.Preresolved) : Syntax /-- `SyntaxNodeKinds` is a set of `SyntaxNodeKind` (implemented as a list). -/ def SyntaxNodeKinds := List SyntaxNodeKind diff --git a/src/Lean/Elab/BuiltinCommand.lean b/src/Lean/Elab/BuiltinCommand.lean index cba007eef0ec..a7b09849a3e0 100644 --- a/src/Lean/Elab/BuiltinCommand.lean +++ b/src/Lean/Elab/BuiltinCommand.lean @@ -120,12 +120,10 @@ private partial def elabChoiceAux (cmds : Array Syntax) (i : Nat) : CommandElabM | Except.error ex => throwError (ex.toMessageData (← getOptions)) @[builtinCommandElab «export»] def elabExport : CommandElab := fun stx => do - -- `stx` is of the form (Command.export "export" "(" (null *) ")") - let id := stx[1].getId - let nss ← resolveNamespace id + let `(export $ns ($ids*)) := stx | throwUnsupportedSyntax + let nss ← resolveNamespace ns let currNamespace ← getCurrNamespace if nss == [currNamespace] then throwError "invalid 'export', self export" - let ids := stx[3].getArgs let mut aliases := #[] for idStx in ids do let id := idStx.getId @@ -133,9 +131,11 @@ private partial def elabChoiceAux (cmds : Array Syntax) (i : Nat) : CommandElabM aliases := aliases.push (currNamespace ++ id, declName) modify fun s => { s with env := aliases.foldl (init := s.env) fun env p => addAlias env p.1 p.2 } -@[builtinCommandElab «open»] def elabOpen : CommandElab := fun n => do - let openDecls ← elabOpenDecl n[1] - modifyScope fun scope => { scope with openDecls := openDecls } +@[builtinCommandElab «open»] def elabOpen : CommandElab + | `(open $decl:openDecl) => do + let openDecls ← elabOpenDecl decl + modifyScope fun scope => { scope with openDecls := openDecls } + | _ => throwUnsupportedSyntax private def typelessBinder? : Syntax → Option (Array (TSyntax [`ident, `Lean.Parser.Term.hole]) × Bool) | `(bracketedBinder|($ids*)) => some <| (ids, true) diff --git a/src/Lean/Elab/BuiltinTerm.lean b/src/Lean/Elab/BuiltinTerm.lean index 6f005ac32eca..7064b5ab0a0c 100644 --- a/src/Lean/Elab/BuiltinTerm.lean +++ b/src/Lean/Elab/BuiltinTerm.lean @@ -268,11 +268,12 @@ private def mkSilentAnnotationIfHole (e : Expr) : TermElabM Expr := do | some msg => elabTermEnsuringType stx[2] expectedType? (errorMsgHeader? := msg) @[builtinTermElab «open»] def elabOpen : TermElab := fun stx expectedType? => do + let `(open $decl in $e) := stx | throwUnsupportedSyntax try pushScope - let openDecls ← elabOpenDecl stx[1] + let openDecls ← elabOpenDecl decl withTheReader Core.Context (fun ctx => { ctx with openDecls := openDecls }) do - elabTerm stx[3] expectedType? + elabTerm e expectedType? finally popScope diff --git a/src/Lean/Elab/Open.lean b/src/Lean/Elab/Open.lean index 9d0576079cbc..196dc4c78621 100644 --- a/src/Lean/Elab/Open.lean +++ b/src/Lean/Elab/Open.lean @@ -13,6 +13,10 @@ variable [Monad m] [STWorld IO.RealWorld m] [MonadEnv m] variable [MonadExceptOf Exception m] [MonadRef m] [AddErrorMessageContext m] variable [AddMessageContext m] [MonadLiftT (ST IO.RealWorld) m] [MonadLog m] +/-- +A local copy of name resolution state that allows us to immediately use new open decls +in further name resolution as in `open Lean Elab`. +-/ structure State where openDecls : List OpenDecl currNamespace : Name @@ -33,19 +37,6 @@ def resolveId (ns : Name) (idStx : Syntax) : M (m := m) Name := do private def addOpenDecl (decl : OpenDecl) : M (m:=m) Unit := modify fun s => { s with openDecls := decl :: s.openDecls } -private def elabOpenSimple (n : Syntax) : M (m:=m) Unit := - -- `open` id+ - for ns in n[0].getArgs do - for ns in (← resolveNamespace ns.getId) do - addOpenDecl (OpenDecl.simple ns []) - activateScoped ns - -private def elabOpenScoped (n : Syntax) : M (m:=m) Unit := - -- `open` `scoped` id+ - for ns in n[1].getArgs do - for ns in (← resolveNamespace ns.getId) do - activateScoped ns - private def resolveNameUsingNamespacesCore (nss : List Name) (idStx : Syntax) : M (m:=m) Name := do let mut exs := #[] let mut result := #[] @@ -66,56 +57,43 @@ private def resolveNameUsingNamespacesCore (nss : List Name) (idStx : Syntax) : else withRef idStx do throwError "ambiguous identifier '{idStx.getId}', possible interpretations: {result.map mkConst}" --- `open` id `(` id+ `)` -private def elabOpenOnly (n : Syntax) : M (m:=m) Unit := do - let nss ← resolveNamespace n[0].getId - for idStx in n[2].getArgs do - let declName ← resolveNameUsingNamespacesCore nss idStx - addOpenDecl (OpenDecl.explicit idStx.getId declName) - --- `open` id `hiding` id+ -private def elabOpenHiding (n : Syntax) : M (m:=m) Unit := do - let ns ← resolveUniqueNamespace n[0].getId - let mut ids : List Name := [] - for idStx in n[2].getArgs do - let _ ← resolveId ns idStx - let id := idStx.getId - ids := id::ids - addOpenDecl (OpenDecl.simple ns ids) - --- `open` id `renaming` sepBy (id `->` id) `,` -private def elabOpenRenaming (n : Syntax) : M (m:=m) Unit := do - let ns ← resolveUniqueNamespace n[0].getId - for stx in n[2].getSepArgs do - let fromStx := stx[0] - let toId := stx[2].getId - let declName ← resolveId ns fromStx - addOpenDecl (OpenDecl.explicit toId declName) - -def elabOpenDecl [MonadResolveName m] (openDeclStx : Syntax) : m (List OpenDecl) := do +def elabOpenDecl [MonadResolveName m] (stx : TSyntax ``Parser.Command.openDecl) : m (List OpenDecl) := do StateRefT'.run' (s := { openDecls := (← getOpenDecls), currNamespace := (← getCurrNamespace) }) do - if openDeclStx.getKind == ``Parser.Command.openSimple then - elabOpenSimple openDeclStx - else if openDeclStx.getKind == ``Parser.Command.openScoped then - elabOpenScoped openDeclStx - else if openDeclStx.getKind == ``Parser.Command.openOnly then - elabOpenOnly openDeclStx - else if openDeclStx.getKind == ``Parser.Command.openHiding then - elabOpenHiding openDeclStx - else - elabOpenRenaming openDeclStx + match stx with + | `(Parser.Command.openDecl| $nss*) => + for ns in nss do + for ns in (← resolveNamespace ns) do + addOpenDecl (OpenDecl.simple ns []) + activateScoped ns + | `(Parser.Command.openDecl| scoped $nss*) => + for ns in nss do + for ns in (← resolveNamespace ns) do + activateScoped ns + | `(Parser.Command.openDecl| $ns ($ids*)) => + let nss ← resolveNamespace ns + for idStx in ids do + let declName ← resolveNameUsingNamespacesCore nss idStx + addOpenDecl (OpenDecl.explicit idStx.getId declName) + | `(Parser.Command.openDecl| $ns hiding $ids*) => + let ns ← resolveUniqueNamespace ns + for id in ids do + let _ ← resolveId ns id + let ids := ids.map (·.getId) |>.toList + addOpenDecl (OpenDecl.simple ns ids) + | `(Parser.Command.openDecl| $ns renaming $[$froms -> $tos],*) => + let ns ← resolveUniqueNamespace ns + for («from», to) in froms.zip tos do + let declName ← resolveId ns «from» + addOpenDecl (OpenDecl.explicit to.getId declName) + | _ => throwUnsupportedSyntax return (← get).openDecls -def resolveOpenDeclId [MonadResolveName m] (ns : Name) (idStx : Syntax) : m Name := do - StateRefT'.run' (s := { openDecls := (← getOpenDecls), currNamespace := (← getCurrNamespace) }) do - OpenDecl.resolveId ns idStx - -def resolveNameUsingNamespaces [MonadResolveName m] (nss : List Name) (idStx : Syntax) : m Name := do +def resolveNameUsingNamespaces [MonadResolveName m] (nss : List Name) (idStx : Ident) : m Name := do StateRefT'.run' (s := { openDecls := (← getOpenDecls), currNamespace := (← getCurrNamespace) }) do resolveNameUsingNamespacesCore nss idStx end OpenDecl -export OpenDecl (elabOpenDecl resolveOpenDeclId resolveNameUsingNamespaces) +export OpenDecl (elabOpenDecl resolveNameUsingNamespaces) end Lean.Elab diff --git a/src/Lean/Elab/Quotation.lean b/src/Lean/Elab/Quotation.lean index 0d1603ff0272..936491465954 100644 --- a/src/Lean/Elab/Quotation.lean +++ b/src/Lean/Elab/Quotation.lean @@ -99,6 +99,11 @@ def tryAddSyntaxNodeKindInfo (stx : Syntax) (k : SyntaxNodeKind) : TermElabM Uni if (← getEnv).contains k then addTermInfo' stx (← mkConstWithFreshMVarLevels k) +instance : Quote Syntax.Preresolved where + quote + | .namespace ns => Unhygienic.run ``(Syntax.Preresolved.namespace $(quote ns)) + | .decl n fs => Unhygienic.run ``(Syntax.Preresolved.decl $(quote n) $(quote fs)) + /-- Elaborate the content of a syntax quotation term -/ private partial def quoteSyntax : Syntax → TermElabM Term | Syntax.ident _ rawVal val preresolved => do @@ -106,11 +111,15 @@ private partial def quoteSyntax : Syntax → TermElabM Term return ← `(Syntax.ident info $(quote rawVal) $(quote val) $(quote preresolved)) -- Add global scopes at compilation time (now), add macro scope at runtime (in the quotation). -- See the paper for details. - let r ← resolveGlobalName val + let consts ← resolveGlobalName val -- extension of the paper algorithm: also store unique section variable names as top-level scopes -- so they can be captured and used inside the section, but not outside - let r' := resolveSectionVariable (← read).sectionVars val - let preresolved := r ++ r' ++ preresolved + let sectionVars := resolveSectionVariable (← read).sectionVars val + -- extension of the paper algorithm: resolve namespaces as well + let namespaces ← resolveNamespaceCore (allowEmpty := true) val + let preresolved := (consts ++ sectionVars).map (fun (n, projs) => Preresolved.decl n projs) ++ + namespaces.map .namespace ++ + preresolved let val := quote val -- `scp` is bound in stxQuot.expand `(Syntax.ident info $(quote rawVal) (addMacroScope mainModule $val scp) $(quote preresolved)) diff --git a/src/Lean/Elab/Tactic/BuiltinTactic.lean b/src/Lean/Elab/Tactic/BuiltinTactic.lean index 46f7556064de..3533457b3716 100644 --- a/src/Lean/Elab/Tactic/BuiltinTactic.lean +++ b/src/Lean/Elab/Tactic/BuiltinTactic.lean @@ -99,11 +99,12 @@ private def getOptRotation (stx : Syntax) : Nat := setGoals <| (← getGoals).rotateRight n @[builtinTactic Parser.Tactic.open] def evalOpen : Tactic := fun stx => do + let `(tactic| open $decl in $tac) := stx | throwUnsupportedSyntax try pushScope - let openDecls ← elabOpenDecl stx[1] + let openDecls ← elabOpenDecl decl withTheReader Core.Context (fun ctx => { ctx with openDecls := openDecls }) do - evalTactic stx[3] + evalTactic tac finally popScope diff --git a/src/Lean/Elab/Term.lean b/src/Lean/Elab/Term.lean index d9bc527a7578..c621a0c03031 100644 --- a/src/Lean/Elab/Term.lean +++ b/src/Lean/Elab/Term.lean @@ -1726,12 +1726,15 @@ private def mkConsts (candidates : List (Name × List String)) (explicitLevels : let const ← mkConst declName explicitLevels return (const, projs) :: result -def resolveName (stx : Syntax) (n : Name) (preresolved : List (Name × List String)) (explicitLevels : List Level) (expectedType? : Option Expr := none) : TermElabM (List (Expr × List String)) := do +def resolveName (stx : Syntax) (n : Name) (preresolved : List Syntax.Preresolved) (explicitLevels : List Level) (expectedType? : Option Expr := none) : TermElabM (List (Expr × List String)) := do addCompletionInfo <| CompletionInfo.id stx stx.getId (danglingDot := false) (← getLCtx) expectedType? if let some (e, projs) ← resolveLocalName n then unless explicitLevels.isEmpty do throwError "invalid use of explicit universe parameters, '{e}' is a local" return [(e, projs)] + let preresolved := preresolved.filterMap fun + | .decl n projs => some (n, projs) + | _ => none -- check for section variable capture by a quotation let ctx ← read if let some (e, projs) := preresolved.findSome? fun (n, projs) => ctx.sectionFVars.find? n |>.map (·, projs) then diff --git a/src/Lean/ResolveName.lean b/src/Lean/ResolveName.lean index abdbe3e026c9..4422e3caa77f 100644 --- a/src/Lean/ResolveName.lean +++ b/src/Lean/ResolveName.lean @@ -204,15 +204,33 @@ instance (m n) [MonadLift m n] [MonadResolveName m] : MonadResolveName n where def resolveGlobalName [Monad m] [MonadResolveName m] [MonadEnv m] (id : Name) : m (List (Name × List String)) := do return ResolveName.resolveGlobalName (← getEnv) (← getCurrNamespace) (← getOpenDecls) id -def resolveNamespace [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Name) : m (List Name) := do - match ResolveName.resolveNamespace (← getEnv) (← getCurrNamespace) (← getOpenDecls) id with - | [] => throwError s!"unknown namespace '{id}'" - | nss => return nss +/-- +Given a namespace name, return a list of possible interpretations. +Names extracted from syntax should be passed to `resolveNamespace` instead. +-/ +def resolveNamespaceCore [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Name) (allowEmpty := false) : m (List Name) := do + let nss := ResolveName.resolveNamespace (← getEnv) (← getCurrNamespace) (← getOpenDecls) id + if !allowEmpty && nss.isEmpty then + throwError s!"unknown namespace '{id}'" + return nss + +/-- Given a namespace identifier, return a list of possible interpretations. -/ +def resolveNamespace [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] : Ident → m (List Name) + | stx@⟨Syntax.ident _ _ n pre⟩ => do + let pre := pre.filterMap fun + | .namespace ns => some ns + | _ => none + if pre.isEmpty then + withRef stx <| resolveNamespaceCore n + else + return pre + | stx => throwErrorAt stx s!"expected identifier" -def resolveUniqueNamespace [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Name) : m Name := do +/-- Given a namespace identifier, return the unique interpretation or else fail. -/ +def resolveUniqueNamespace [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Ident) : m Name := do match (← resolveNamespace id) with | [ns] => return ns - | nss => throwError s!"ambiguous namespace '{id}', possible interpretations: '{nss}'" + | nss => throwError s!"ambiguous namespace '{id.getId}', possible interpretations: '{nss}'" /-- Given a name `n`, return a list of possible interpretations for global constants. @@ -254,7 +272,9 @@ After `open Foo open Boo`, we have -/ def resolveGlobalConst [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] : Syntax → m (List Name) | stx@(Syntax.ident _ _ n pre) => do - let pre := pre.filterMap fun (n, fields) => if fields.isEmpty then some n else none + let pre := pre.filterMap fun + | .decl n [] => some n + | _ => none if pre.isEmpty then withRef stx <| resolveGlobalConstCore n else diff --git a/stage0/src/Init/Prelude.lean b/stage0/src/Init/Prelude.lean index 1e77419d6e2f..deff842eb6d2 100644 --- a/stage0/src/Init/Prelude.lean +++ b/stage0/src/Init/Prelude.lean @@ -83,7 +83,7 @@ example : foo.default = (default, default) := rfl ``` -/ -@[reducible] def inferInstance {α : Sort u} [i : α] : α := i +abbrev inferInstance {α : Sort u} [i : α] : α := i set_option checkBinderAnnotations false in /-- `inferInstanceAs α` synthesizes a value of any target type by typeclass @@ -97,7 +97,7 @@ does.) Example: #check inferInstanceAs (Inhabited Nat) -- Inhabited Nat ``` -/ -@[reducible] def inferInstanceAs (α : Sort u) [i : α] : α := i +abbrev inferInstanceAs (α : Sort u) [i : α] : α := i set_option bootstrap.inductiveCheckResultingUniverse false in /-- @@ -3399,6 +3399,17 @@ abbrev SyntaxNodeKind := Name /-! # Syntax AST -/ +/-- +Binding information resolved and stored at compile time of a syntax quotation. +Note: We do not statically know whether a syntax expects a namespace or term name, +so a `Syntax.ident` may contain both preresolution kinds. +-/ +inductive Syntax.Preresolved where + | /-- A potential namespace reference -/ + namespace (ns : Name) + | /-- A potential global constant or section variable reference, with additional field accesses -/ + decl (n : Name) (fields : List String) + /-- Syntax objects used by the parser, macro expander, delaborator, etc. -/ diff --git a/stage0/src/Lean/Compiler/CSE.lean b/stage0/src/Lean/Compiler/CSE.lean index 785fa5fe30bd..8c2e429b4f34 100644 --- a/stage0/src/Lean/Compiler/CSE.lean +++ b/stage0/src/Lean/Compiler/CSE.lean @@ -27,13 +27,13 @@ partial def visitCases (casesInfo : CasesInfo) (cases : Expr) : M Expr := do partial def visitLambda (e : Expr) : M Expr := withNewScope do - let (as, e) ← Compiler.visitLambda e - let e ← mkLetUsingScope (← visitLet e) + let (as, e) ← Compiler.visitLambdaCore e + let e ← mkLetUsingScope (← visitLet e as) mkLambda as e -partial def visitLet (e : Expr) : M Expr := do +partial def visitLet (e : Expr) (xs : Array Expr): M Expr := do let saved ← get - try go e #[] finally set saved + try go e xs finally set saved where go (e : Expr) (xs : Array Expr) : M Expr := do match e with diff --git a/stage0/src/Lean/Compiler/Check.lean b/stage0/src/Lean/Compiler/Check.lean index 85533b78c063..3c13a96927ac 100644 --- a/stage0/src/Lean/Compiler/Check.lean +++ b/stage0/src/Lean/Compiler/Check.lean @@ -14,6 +14,7 @@ Type checker for LCNF expressions -/ structure Check.Config where + checkJpScope : Bool := true terminalCasesOnly : Bool := true structure Check.Context where @@ -69,7 +70,7 @@ where checkBlock b (xs.push x) | _ => if let some fvarId ← isJump? e then - unless (← read).jps.contains fvarId do + unless !cfg.checkJpScope || (← read).jps.contains fvarId do /- We cannot jump to join points defined out of the scope of a local function declaration. For example, the following is an invalid LCNF. diff --git a/stage0/src/Lean/Compiler/CompilerM.lean b/stage0/src/Lean/Compiler/CompilerM.lean index ff98843a3d3b..3e317cac4163 100644 --- a/stage0/src/Lean/Compiler/CompilerM.lean +++ b/stage0/src/Lean/Compiler/CompilerM.lean @@ -21,6 +21,7 @@ structure CompilerM.State where letFVars : Array Expr := #[] /-- Next auxiliary variable suffix -/ nextIdx : Nat := 1 +deriving Inhabited abbrev CompilerM := StateRefT CompilerM.State CoreM @@ -56,6 +57,12 @@ def mkLetDecl (binderName : Name) (type : Expr) (value : Expr) (nonDep : Bool) : modify fun s => { s with lctx := s.lctx.mkLetDecl fvarId binderName type value nonDep, letFVars := s.letFVars.push x } return x +def mkFreshLetVarIdx : CompilerM Nat := do + modifyGet fun s => (s.nextIdx, { s with nextIdx := s.nextIdx +1 }) + +def mkAuxLetDeclName (prefixName := `_x) : CompilerM Name := + return .num prefixName (← mkFreshLetVarIdx) + /-- Create a new auxiliary let declaration with value `e` The name of the declaration is guaranteed to be unique. @@ -65,10 +72,7 @@ def mkAuxLetDecl (e : Expr) (prefixName := `_x) : CompilerM Expr := do if e.isFVar then return e else - try - mkLetDecl (.num prefixName (← get).nextIdx) (← inferType e) e (nonDep := false) - finally - modify fun s => { s with nextIdx := s.nextIdx + 1 } + mkLetDecl (← mkAuxLetDeclName prefixName) (← inferType e) e (nonDep := false) /-- Create an auxiliary let declaration with value `e`, that is a join point. @@ -97,7 +101,7 @@ let-declarations that are safe to unfold without producing code blowup, and join Remark: user-facing names provided by users are preserved. We keep them as the prefix of the new unique names. -/ -def ensureUniqueLetVarNames (e : Expr) : CoreM Expr := +def ensureUniqueLetVarNamesCore (e : Expr) : StateRefT Nat CoreM Expr := let pre (e : Expr) : StateRefT Nat CoreM TransformStep := do match e with | .letE binderName type value body nonDep => @@ -107,17 +111,25 @@ def ensureUniqueLetVarNames (e : Expr) : CoreM Expr := | _ => .num binderName idx return .visit <| .letE binderName' type value body nonDep | _ => return .visit e - Core.transform e pre |>.run' 1 + Core.transform e pre + +def ensureUniqueLetVarNames (e : Expr) : CompilerM Expr := do + let (e, nextIdx) ← ensureUniqueLetVarNamesCore e |>.run (← get).nextIdx + modify fun s => { s with nextIdx } + return e /-- Move through all consecutive lambda abstractions at the top level of `e`. -Returning the body of the last one we find with all bound variables -instantiated as new free variables. +Returning the body of the last one we find with **loose bound variables**. Returns a tuple consisting of: -1. An `Array` of all the newly created free variables -2. The (fully instantiated) body of the last lambda binder that was visited +1. An `Array` of all the newly created free variables. +2. The body of the last lambda binder that was visited. + The caller is responsible for replacing the loose bound variables + with the newly created free variables. + +See `visitLambda`. -/ -def visitLambda (e : Expr) : CompilerM (Array Expr × Expr) := +def visitLambdaCore (e : Expr) : CompilerM (Array Expr × Expr) := go e #[] where go (e : Expr) (fvars : Array Expr) := do @@ -126,7 +138,19 @@ where let fvar ← mkLocalDecl binderName type binderInfo go body (fvars.push fvar) else - return (fvars, e.instantiateRev fvars) + return (fvars, e) + +/-- +Move through all consecutive lambda abstractions at the top level of `e`. +Returning the body of the last one we find with all bound variables +instantiated as new free variables. +Returns a tuple consisting of: +1. An `Array` of all the newly created free variables +2. The (fully instantiated) body of the last lambda binder that was visited +-/ +def visitLambda (e : Expr) : CompilerM (Array Expr × Expr) := do + let (fvars, e) ← visitLambdaCore e + return (fvars, e.instantiateRev fvars) /-- Given an expression representing a `match` return a tuple consisting of: @@ -147,7 +171,7 @@ def visitMatch (cases : Expr) (casesInfo : CasesInfo) : CompilerM (Expr × Array arms := arms.push (←visitLambda args[i]!).snd return (motive, discrs, arms) -def withNewScopeImp (x : CompilerM α) : CompilerM α := do +@[inline] def withNewScopeImp (x : CompilerM α) : CompilerM α := do let saved ← get modify fun s => { s with letFVars := #[] } try x @@ -155,7 +179,7 @@ def withNewScopeImp (x : CompilerM α) : CompilerM α := do let saved := { saved with nextIdx := (← get).nextIdx } set saved -def withNewScope [MonadFunctorT CompilerM m] (x : m α) : m α := +@[inline] def withNewScope [MonadFunctorT CompilerM m] (x : m α) : m α := monadMap (m := CompilerM) withNewScopeImp x /-- @@ -175,10 +199,10 @@ class VisitLet (m : Type → Type) where export VisitLet (visitLet) -def visitLetImp (e : Expr) (f : Name → Expr → CompilerM Expr) : CompilerM Expr := +@[inline] def visitLetImp (e : Expr) (f : Name → Expr → CompilerM Expr) : CompilerM Expr := go e #[] where - go (e : Expr) (fvars : Array Expr) : CompilerM Expr := do + @[specialize] go (e : Expr) (fvars : Array Expr) : CompilerM Expr := do if let .letE binderName type value body nonDep := e then let type := type.instantiateRev fvars let value := value.instantiateRev fvars @@ -284,8 +308,8 @@ partial def attachJp (e : Expr) (jp : Expr) : CompilerM Expr := do where visitLambda (e : Expr) : ReaderT FVarIdSet CompilerM Expr := do withNewScope do - let (as, e) ← Compiler.visitLambda e - let e ← mkLetUsingScope (← visitLet e #[]) + let (as, e) ← Compiler.visitLambdaCore e + let e ← mkLetUsingScope (← visitLet e as) mkLambda as e visitCases (casesInfo : CasesInfo) (cases : Expr) : ReaderT FVarIdSet CompilerM Expr := do @@ -299,13 +323,19 @@ where visitLet (e : Expr) (xs : Array Expr) : ReaderT FVarIdSet CompilerM Expr := do match e with | .letE binderName type value body nonDep => + let mkDecl (type value : Expr) := do + let x ← mkLetDecl binderName type value nonDep + withReader (fun jps => if isJpBinderName binderName then jps.insert x.fvarId! else jps) do + visitLet body (xs.push x) let type := type.instantiateRev xs - let mut value := value.instantiateRev xs + let value := value.instantiateRev xs if isJpBinderName binderName then - value ← visitLambda value - let x ← mkLetDecl binderName type value nonDep - withReader (fun jps => if isJpBinderName binderName then jps.insert x.fvarId! else jps) do - visitLet body (xs.push x) + let value ← visitLambda value + -- Recall that the resulting type of join point may change after the attachment + let type ← inferType value + mkDecl type value + else + mkDecl type value | _ => let e := e.instantiateRev xs if let some fvarId ← isJump? e then diff --git a/stage0/src/Lean/Compiler/Decl.lean b/stage0/src/Lean/Compiler/Decl.lean index a1e592878aef..17031d3d39c5 100644 --- a/stage0/src/Lean/Compiler/Decl.lean +++ b/stage0/src/Lean/Compiler/Decl.lean @@ -166,7 +166,7 @@ Make sure all let-declarations have unique user-facing names. See `Compiler.ensureUniqueLetVarNames` -/ def Decl.ensureUniqueLetVarNames (decl : Decl) : CoreM Decl := do - return { decl with value := (← Compiler.ensureUniqueLetVarNames decl.value) } + return { decl with value := (← ensureUniqueLetVarNamesCore decl.value |>.run' 1) } def Decl.getArity (decl : Decl) : Nat := getLambdaArity decl.value diff --git a/stage0/src/Lean/Compiler/InferType.lean b/stage0/src/Lean/Compiler/InferType.lean index d22ec8108e31..55ded6f0dec3 100644 --- a/stage0/src/Lean/Compiler/InferType.lean +++ b/stage0/src/Lean/Compiler/InferType.lean @@ -69,11 +69,12 @@ mutual match fType with | .forallE _ _ b _ => fType := b | _ => - match fType.instantiateRevRange j i args |>.headBeta with + fType := fType.instantiateRevRange j i args |>.headBeta + match fType with | .forallE _ _ b _ => j := i; fType := b | _ => if fType.isAnyType then return anyTypeExpr - throwError "function expected{indentExpr f}" + throwError "function expected{indentExpr (mkAppN f args[:i])} : {fType}\nfunction type{indentExpr (← inferType f)}" return fType.instantiateRevRange j args.size args |>.headBeta partial def inferProjType (structName : Name) (idx : Nat) (s : Expr) : InferTypeM Expr := do diff --git a/stage0/src/Lean/Compiler/JoinPoints.lean b/stage0/src/Lean/Compiler/JoinPoints.lean index 3baaad9b32d1..c74cd75b2023 100644 --- a/stage0/src/Lean/Compiler/JoinPoints.lean +++ b/stage0/src/Lean/Compiler/JoinPoints.lean @@ -74,14 +74,44 @@ end Visitors namespace JoinPointFinder -abbrev M := StateRefT (Std.HashMap Name Nat) CompilerM +structure CandidateInfo where + arity : Nat + associated : Std.HashSet Name + deriving Inhabited + +abbrev M := ReaderT (Option Name) StateRefT (Std.HashMap Name CandidateInfo) CompilerM + +private def findCandidate? (name : Name) : M (Option CandidateInfo) := do + return (← get).find? name + +private partial def eraseCandidate (name : Name) : M Unit := do + if let some info ← findCandidate? name then + modify (fun candidates => candidates.erase name) + info.associated.forM eraseCandidate private partial def removeCandidatesContainedIn (e : Expr) : M Unit := do let remover := fun fvarId => do let some decl ← findDecl? fvarId | unreachable! - modify (fun jps? => jps?.erase decl.userName) + eraseCandidate decl.userName forEachFVar e remover +private def addCandidate (name : Name) (arity : Nat) : M Unit := do + let cinfo := { arity, associated := .empty } + modify (fun cs => cs.insert name cinfo) + +private def addDependency (src : Name) (target : Name) : M Unit := do + if let some targetInfo ← findCandidate? target then + modify (fun cs => cs.insert target { targetInfo with associated := targetInfo.associated.insert src }) + else + eraseCandidate src + +def declIsInScope (decl : LocalDecl) : CompilerM Bool := do + let fvars := (←getThe CompilerM.State).letFVars + fvars.anyM (fun fvar => do + let scopeDecl := (←findDecl? fvar.fvarId!).get! + return scopeDecl.userName == decl.userName + ) + /-- Return a set of let declaration names inside of `e` that qualify as a join point that is: @@ -92,25 +122,28 @@ Since declaration names are unique inside of LCNF the let bindings and call sites can be uniquely identified by this. -/ partial def findJoinPoints (e : Expr) : CompilerM (Array Name) := do - let (_, names) ← visitLambda goTailApp removeCandidatesContainedIn goLetValue e |>.run .empty - return names.toArray.map Prod.fst + let (_, state) ← visitLambda goTailApp removeCandidatesContainedIn goLetValue e |>.run none |>.run .empty + return state.toArray.map Prod.fst where goLetValue (userName : Name) (value : Expr) : M Expr := do - if let .lam .. := value then - withNewScope do - let (vars, body) ← Compiler.visitLambda value - modify (fun jps? => jps?.insert userName vars.size) + match value with + | .lam .. => withNewScope do + let (vars, body) ← Compiler.visitLambda value + addCandidate userName vars.size + withReader (fun _ => some userName) do visitTails goTailApp removeCandidatesContainedIn goLetValue body - else - visitTails goTailApp removeCandidatesContainedIn goLetValue value + | _ => removeCandidatesContainedIn value return value goTailApp (fvarId : FVarId) (e : Expr) : M Unit := do let some decl ← findDecl? fvarId | unreachable! - if let some jpArity := (←get).find? decl.userName then + if let some candidateInfo ← findCandidate? decl.userName then let args := e.getAppNumArgs - if args != jpArity then - modify (fun jps? => jps?.erase decl.userName) + if args != candidateInfo.arity then + eraseCandidate decl.userName + else if let some upperCandidate ← read then + if !(←declIsInScope decl) then + addDependency decl.userName upperCandidate end JoinPointFinder @@ -120,7 +153,6 @@ namespace JoinPointChecker Throws an exception if `e` contains a join point. -/ def containsNoJp (e : Expr) : CompilerM Unit := do - trace[Compiler.step] s!"Checking whether {e} contains jp" let checker := fun fvarId => do let some decl ← findDecl? fvarId | unreachable! if decl.isJp then diff --git a/stage0/src/Lean/Compiler/LCNF.lean b/stage0/src/Lean/Compiler/LCNF.lean index bcc9da5beef5..7d41e67499d7 100644 --- a/stage0/src/Lean/Compiler/LCNF.lean +++ b/stage0/src/Lean/Compiler/LCNF.lean @@ -41,14 +41,16 @@ structure Context where structure State where /-- Local context containing the original Lean types (not LCNF ones). -/ - lctx : LocalContext := {} + lctx : LocalContext := {} /-- Local context containing Lean LCNF types. -/ lctx' : LocalContext := {} letFVars : Array Expr := #[] /-- Next auxiliary variable suffix -/ nextIdx : Nat := 1 /-- Cache from Lean regular expression to LCNF expression. -/ - cache : PersistentExprMap Expr := {} + cache : Std.PHashMap (Expr × Bool) Expr := {} + /-- `toLCNFType` cache -/ + typeCache : Std.PHashMap Expr Expr := {} abbrev M := ReaderT Context $ StateRefT State CoreM @@ -67,13 +69,21 @@ def withNewRootScope (x : M α) : M α := do try withRoot true x finally - let saved := { saved with nextIdx := (← get).nextIdx } + let saved := { saved with nextIdx := (← get).nextIdx, typeCache := (← get).typeCache } set saved +def toLCNFType (type : Expr) : M Expr := do + match (← get).typeCache.find? type with + | some type' => return type' + | none => + let type' ← liftMetaM <| Compiler.toLCNFType type + modify fun s => { s with typeCache := s.typeCache.insert type type' } + return type' + /-- Create a new local declaration using a Lean regular type. -/ def mkLocalDecl (binderName : Name) (type : Expr) (bi := BinderInfo.default) : M Expr := do let fvarId ← mkFreshFVarId - let type' ← liftMetaM <| toLCNFType type + let type' ← toLCNFType type modify fun s => { s with lctx := s.lctx.mkLocalDecl fvarId binderName type bi lctx' := s.lctx'.mkLocalDecl fvarId binderName type' bi @@ -170,7 +180,8 @@ partial def toLCNF (e : Expr) : CoreM Expr := do return s.lctx'.mkLambda s.letFVars e where visitCore (e : Expr) : M Expr := withIncRecDepth do - if let some e := (← get).cache.find? e then + let key := (e, (← read).root) + if let some e := (← get).cache.find? key then return e let r ← match e with | .app .. => visitApp e @@ -184,7 +195,7 @@ where | .mvar .. => throwError "unexpected occurrence of metavariable in code generator{indentExpr e}" | .bvar .. => unreachable! | _ => pure e - modify fun s => { s with cache := s.cache.insert e r } + modify fun s => { s with cache := s.cache.insert key r } return r visit (e : Expr) : M Expr := withIncRecDepth do @@ -217,7 +228,7 @@ where return mkConst ``lcErased if (← liftMetaM <| Meta.isTypeFormerType type) then /- Types and Type formers are not put into A-normal form -/ - return (← liftMetaM <| toLCNFType e) + toLCNFType e else withRoot false <| visitCore e @@ -273,7 +284,7 @@ where visitCases (casesInfo : CasesInfo) (e : Expr) : M Expr := etaIfUnderApplied e casesInfo.arity do let mut args := e.getAppArgs - let mut resultType ← liftMetaM do toLCNFType (← Meta.inferType (mkAppN e.getAppFn args[:casesInfo.arity])) + let mut resultType ← toLCNFType (← liftMetaM do Meta.inferType (mkAppN e.getAppFn args[:casesInfo.arity])) let mut discrTypes := #[] for i in [:casesInfo.numParams] do -- `cases` and `match` parameters are irrelevant at LCNF @@ -318,7 +329,7 @@ where etaIfUnderApplied e arity do let args := e.getAppArgs let f := e.getAppFn - let recType ← liftMetaM do toLCNFType (← Meta.inferType (mkAppN f args[:arity])) + let recType ← toLCNFType (← liftMetaM do Meta.inferType (mkAppN f args[:arity])) let minor := if e.isAppOf ``Eq.rec || e.isAppOf ``Eq.ndrec then args[3]! else args[5]! let minor ← visit minor let minorType ← inferType minor @@ -333,7 +344,7 @@ where visitFalseRec (e : Expr) : M Expr := let arity := 2 etaIfUnderApplied e arity do - let type ← liftMetaM do toLCNFType (← Meta.inferType e) + let type ← toLCNFType (← liftMetaM do Meta.inferType e) mkAuxLetDecl (← mkLcUnreachable type) visitAndRec (e : Expr) : M Expr := @@ -444,7 +455,7 @@ where if (← liftMetaM <| Meta.isProp type <||> Meta.isTypeFormerType type) then visitLet body (xs.push value) else - let type' ← liftMetaM <| toLCNFType type + let type' ← toLCNFType type let value' ← withRoot true <| visit value let x ← mkLetDecl binderName type value type' value' visitLet body (xs.push x) diff --git a/stage0/src/Lean/Compiler/LCNFTypes.lean b/stage0/src/Lean/Compiler/LCNFTypes.lean index de57f71b1aa4..57f05ff58839 100644 --- a/stage0/src/Lean/Compiler/LCNFTypes.lean +++ b/stage0/src/Lean/Compiler/LCNFTypes.lean @@ -93,18 +93,25 @@ partial def toLCNFType (type : Expr) : MetaM Expr := do return b else return (Expr.lam n d (b.abstract #[x]) bi).eta - | .forallE n d b bi => - withLocalDecl n bi d fun x => do - let borrowed := isMarkedBorrowed d - let mut d ← toLCNFType d - if borrowed then - d := markBorrowed d - let b ← toLCNFType (b.instantiate1 x) - return .forallE n d (b.abstract #[x]) bi + | .forallE .. => visitForall type #[] | .app .. => type.withApp visitApp | .fvar .. => visitApp type #[] | _ => return anyTypeExpr where + visitForall (e : Expr) (xs : Array Expr) : MetaM Expr := do + match e with + | .forallE n d b bi => + let d := d.instantiateRev xs + withLocalDecl n bi d fun x => do + let borrowed := isMarkedBorrowed d + let mut d := (← toLCNFType d).abstract xs + if borrowed then + d := markBorrowed d + return .forallE n d (← visitForall b (xs.push x)) bi + | _ => + let e ← toLCNFType (e.instantiateRev xs) + return e.abstract xs + visitApp (f : Expr) (args : Array Expr) := do let fNew ← match f with | .const declName us => @@ -185,4 +192,14 @@ def isTypeFormerType (type : Expr) : Bool := | .forallE _ _ b _ => isTypeFormerType b | _ => type.isAnyType +/-- +`isClass? type` return `some ClsName` if the LCNF `type` is an instance of the class `ClsName`. +-/ +def isClass? (type : Expr) : CoreM (Option Name) := do + let .const declName _ := type.getAppFn | return none + if isClass (← getEnv) declName then + return declName + else + return none + end Lean.Compiler \ No newline at end of file diff --git a/stage0/src/Lean/Compiler/Main.lean b/stage0/src/Lean/Compiler/Main.lean index 03143bbb747c..229485e39bda 100644 --- a/stage0/src/Lean/Compiler/Main.lean +++ b/stage0/src/Lean/Compiler/Main.lean @@ -8,6 +8,7 @@ import Lean.Compiler.TerminalCases import Lean.Compiler.CSE import Lean.Compiler.Stage1 import Lean.Compiler.Simp +import Lean.Compiler.PullLocalDecls namespace Lean.Compiler @@ -53,12 +54,17 @@ def compileStage1Impl (declNames : Array Name) : CoreM (Array Decl) := do checkpoint `init decls { terminalCasesOnly := false } let decls ← decls.mapM (·.terminalCases) checkpoint `terminalCases decls - let decls ← decls.mapM (·.simp) - checkpoint `simp decls - -- Remark: add simplification step here, `cse` is useful after simplification + decls.forM fun decl => do trace[Compiler.jp] "{decl.name}: {(← JoinPoints.JoinPointFinder.findJoinPoints decl.value |>.run' {})}" + let decls ← decls.mapM (·.pullInstances) + checkpoint `pullInstances decls let decls ← decls.mapM (·.cse) checkpoint `cse decls + let decls ← decls.mapM (·.simp) + checkpoint `simp decls + -- let decls ← decls.mapM (·.cse) + -- checkpoint `cse decls saveStage1Decls decls + decls.forM fun decl => do trace[Compiler.stat] "{decl.name}: {← getLCNFSize decl.value}" return decls /-- @@ -70,9 +76,12 @@ def compile (declNames : Array Name) : CoreM Unit := do profileitM Exception "co builtin_initialize registerTraceClass `Compiler + registerTraceClass `Compiler.stat registerTraceClass `Compiler.init (inherited := true) registerTraceClass `Compiler.terminalCases (inherited := true) registerTraceClass `Compiler.simp (inherited := true) + registerTraceClass `Compiler.pullInstances (inherited := true) registerTraceClass `Compiler.cse (inherited := true) + registerTraceClass `Compiler.jp end Lean.Compiler diff --git a/stage0/src/Lean/Compiler/PullLocalDecls.lean b/stage0/src/Lean/Compiler/PullLocalDecls.lean new file mode 100644 index 000000000000..a9137e326793 --- /dev/null +++ b/stage0/src/Lean/Compiler/PullLocalDecls.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2022 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Lean.Compiler.CompilerM +import Lean.Compiler.Decl + +namespace Lean.Compiler +namespace PullLocalDecls + +structure Context where + isCandidateFn : LocalDecl → FVarIdSet → CompilerM Bool + +structure State where + toPull : FVarIdSet := {} + +abbrev PullM := ReaderT Context $ StateRefT State CompilerM + +def dependsOn (e : Expr) (fvarIdSet : FVarIdSet) : Bool := + Option.isSome <| e.find? fun e => e.isFVar && fvarIdSet.contains e.fvarId! + +def mkLambda (xs : Array Expr) (letFVarsSavedSize : Nat) (e : Expr) : PullM Expr := do + let letFVars := (← getThe CompilerM.State).letFVars + let toPull := (← get).toPull + let mut xs := xs + let mut included : FVarIdSet := xs.foldl (init := {}) fun s x => s.insert x.fvarId! + let mut keep := #[] -- Variables to keep pulling + for i in [letFVarsSavedSize : letFVars.size] do + let y := letFVars[i]! + let some (.ldecl _ fvarId _ type value _) ← findDecl? y.fvarId! | unreachable! + if toPull.contains fvarId then + if !dependsOn value included && !dependsOn type included then + keep := keep.push y + continue + xs := xs.push y + included := included.insert fvarId + modifyThe CompilerM.State fun s => { s with letFVars := s.letFVars.shrink letFVarsSavedSize ++ keep } + -- Remark: we could remove all variables in `xs` from the local context + Compiler.mkLambda xs e + +mutual + +partial def visitLambda (e : Expr) (topLevel := false) : PullM Expr := do + let (as, e) ← Compiler.visitLambdaCore e + let letFVarsSavedSize := (← getThe CompilerM.State).letFVars.size + let e ← visitLet e as + if topLevel then + Compiler.mkLambda as (← mkLetUsingScope e) + else + mkLambda as letFVarsSavedSize e + +partial def visitCases (casesInfo : CasesInfo) (cases : Expr) : PullM Expr := do + let mut args := cases.getAppArgs + for i in casesInfo.altsRange do + args ← args.modifyM i visitLambda + return mkAppN cases.getAppFn args + +partial def visitLet (e : Expr) (xs : Array Expr := #[]) : PullM Expr := do + match e with + | .letE binderName type value body nonDep => + let type := type.instantiateRev xs + let mut value := value.instantiateRev xs + if value.isLambda then + value ← visitLambda value + let x ← mkLetDecl binderName type value nonDep + let some localDecl ← findDecl? x.fvarId! | unreachable! + if (← (← read).isCandidateFn localDecl (← get).toPull) then + trace[Compiler.pullLocalDecls.candidate] "{x}" + modify fun s => { s with toPull := s.toPull.insert localDecl.fvarId } + visitLet body (xs.push x) + | e => + let e := e.instantiateRev xs + if let some casesInfo ← isCasesApp? e then + visitCases casesInfo e + else + return e + +end + +end PullLocalDecls + +def Decl.pullLocalDecls (decl : Decl) (isCandidateFn : LocalDecl → FVarIdSet → CompilerM Bool) : CoreM Decl := + decl.mapValue fun value => PullLocalDecls.visitLambda value (topLevel := true) |>.run { isCandidateFn } |>.run' {} + +def Decl.pullInstances (decl : Decl) : CoreM Decl := + decl.pullLocalDecls fun localDecl candidates => do + if (← isClass? localDecl.type).isSome then + return true + else if let .proj _ _ (.fvar fvarId) := localDecl.value then + return candidates.contains fvarId + else + return false + +builtin_initialize + registerTraceClass `Compiler.pullLocalDecls.candidate + +end Lean.Compiler diff --git a/stage0/src/Lean/Compiler/Simp.lean b/stage0/src/Lean/Compiler/Simp.lean index 9869c1c8a412..4d602be188a1 100644 --- a/stage0/src/Lean/Compiler/Simp.lean +++ b/stage0/src/Lean/Compiler/Simp.lean @@ -27,146 +27,381 @@ partial def findExpr (e : Expr) (skipMData := true): CompilerM Expr := do | .mdata _ e' => if skipMData then findExpr e' else return e | _ => return e +inductive Occ where + | once + | many + deriving Repr, Inhabited + /-- Local function declaration statistics. Remark: we use the `userName` as the key. Thus, `ensureUniqueLetVarNames` must be used before collectin stastistics. -/ -structure InlineStats where - /-- - Mapping from local function name to the number of times it is used - in a declaration. - -/ - numOccs : Std.HashMap Name Nat := {} +structure OccInfo where /-- - Mapping from local function name to their LCNF size. + Mapping from local function name to occurrence information. -/ - size : Std.HashMap Name Nat := {} + map : Std.HashMap Name Occ := {} + deriving Inhabited -def InlineStats.format (s : InlineStats) : Format := Id.run do +def OccInfo.format (s : OccInfo) : Format := Id.run do let mut result := Format.nil - for (k, n) in s.numOccs.toList do - let some size := s.size.find? k | pure () - result := result ++ "\n" ++ f!"{k} ↦ {n}, {size}" - pure () + for (k, n) in s.map.toList do + result := result ++ "\n" ++ f!"{k} ↦ {repr n}" return result -def InlineStats.shouldInline (s : InlineStats) (k : Name) : Bool := Id.run do - let some numOccs := s.numOccs.find? k | return false - if numOccs == 1 then return true - let some sz := s.size.find? k | return false - return sz == 1 +instance : ToFormat OccInfo where + format := OccInfo.format + +def OccInfo.add (s : OccInfo) (key : Name) : OccInfo := + match s with + | { map } => + match map.find? key with + | some .once => { map := map.insert key .many } + | none => { map := map.insert key .once } + | _ => { map } + +structure Config where + smallThreshold : Nat := 1 -instance : ToFormat InlineStats where - format := InlineStats.format +structure Context where + config : Config := {} -partial def collectInlineStats (e : Expr) : CoreM InlineStats := do - let ((_, s), _) ← goLambda e |>.run {} |>.run {} - return s +structure State where + /-- + (Approximate) occurence information for local function declarations. + -/ + occInfo : OccInfo := {} + simplified : Bool := false + deriving Inhabited + +abbrev SimpM := ReaderT Context $ StateRefT State CompilerM + +/-- Ensure binder names are unique, and update occurrence information -/ +partial def internalize (e : Expr) : SimpM Expr := do + visitLambda e where - goLambda (e : Expr) : StateRefT InlineStats CompilerM Unit := do + visitLambda (e : Expr) : SimpM Expr := do withNewScope do - let (_, body) ← visitLambda e - go body + let (as, e) ← Compiler.visitLambdaCore e + let e ← mkLetUsingScope (← visitLet e as) + mkLambda as e - goValue (value : Expr) : StateRefT InlineStats CompilerM Unit := do - match value with - | .lam .. => goLambda value - | .app .. => - match (← findLambda? value.getAppFn) with + visitCases (casesInfo : CasesInfo) (cases : Expr) : SimpM Expr := do + let mut args := cases.getAppArgs + for i in casesInfo.altsRange do + args ← args.modifyM i visitLambda + return mkAppN cases.getAppFn args + + visitValue (e : Expr) : SimpM Unit := do + if e.isApp then + match (← findLambda? e.getAppFn) with | some localDecl => - trace[Meta.debug] "found decl {localDecl.userName}" if localDecl.value.isLambda then let key := localDecl.userName - match (← get).numOccs.find? localDecl.userName with - | some numOccs => modify fun s => { s with numOccs := s.numOccs.insert key (numOccs + 1) } - | _ => - let sz ← getLCNFSize localDecl.value - modify fun { numOccs, size } => { numOccs := numOccs.insert key 1, size := size.insert key sz } + modify fun s => { s with occInfo := s.occInfo.add key } | _ => pure () - | _ => pure () - go (e : Expr) : StateRefT InlineStats CompilerM Unit := do + visitLet (e : Expr) (xs : Array Expr) : SimpM Expr := do match e with - | .letE .. => - withNewScope do - let body ← visitLet e fun _ value => do goValue value; return value - go body - | e => + | .letE binderName type value body nonDep => + let idx ← mkFreshLetVarIdx + let binderName' := match binderName with + | .num p _ => .num p idx + | _ => .num binderName idx + let type := type.instantiateRev xs + let mut value := value.instantiateRev xs + if value.isLambda then + value ← visitLambda value + else + visitValue value + let x ← mkLetDecl binderName' type value nonDep + visitLet body (xs.push x) + | _ => + let e := e.instantiateRev xs if let some casesInfo ← isCasesApp? e then - let args := e.getAppArgs - for i in casesInfo.altsRange do - goLambda args[i]! + visitCases casesInfo e else - goValue e + visitValue e + return e -structure Config where - increaseFactor : Nat := 2 +def markSimplified : SimpM Unit := + modify fun s => { s with simplified := true } -structure Context where - config : Config := {} - /-- - Statistics for deciding whether to inline local function declarations. - -/ - stats : InlineStats - /-- - We only inline local declarations when `localInline` is `true`. - We set it to `false` when we are inlining a non local definition - that may have let-declarations whose names collide with the ones - stored at `stats`. - -/ - localInline : Bool := true +def findCtor (e : Expr) : SimpM Expr := do + -- TODO: add support for mapping discriminants to constructors in branches + findExpr e -structure State where - simplified : Bool := false +/-- +Try to simplify projections `.proj _ i s` where `s` is constructor. +-/ +def simpProj? (e : Expr) : OptionT SimpM Expr := do + let .proj _ i s := e | failure + let s ← findCtor s + let some (ctorVal, args) := s.constructorApp? (← getEnv) | failure + markSimplified + return args[ctorVal.numParams + i]! -abbrev SimpM := ReaderT Context $ StateRefT State CompilerM +/-- +Application over application. +``` +let _x.i := f a +_x.i b +``` +is simplified to `f a b`. +-/ +def simpAppApp? (e : Expr) : OptionT SimpM Expr := do + guard e.isApp + let f := e.getAppFn + guard f.isFVar + let f ← findExpr f + guard <| f.isApp || f.isConst + markSimplified + return mkAppN f e.getAppArgs -def markSimplified : SimpM Unit := - modify fun s => { s with simplified := true } +def shouldInlineLocal (localDecl : LocalDecl) : SimpM Bool := do + match (← get).occInfo.map.find? localDecl.userName with + | some .once => return true + | _ => lcnfSizeLe localDecl.value (← read).config.smallThreshold -def shouldInline (localDecl : LocalDecl) : SimpM Bool := - return (← read).localInline && (← read).stats.shouldInline localDecl.userName +structure InlineCandidateInfo where + isLocal : Bool + arity : Nat + /-- Value (lambda expression) of the function to be inlined. -/ + value : Expr -def inlineCandidate? (e : Expr) : SimpM (Option Nat) := do +def inlineCandidate? (e : Expr) : SimpM (Option InlineCandidateInfo) := do let f := e.getAppFn - let arity ← match f with - | .const declName _ => - unless hasInlineAttribute (← getEnv) declName do return none - -- TODO: check whether function is recursive or not. - -- We can skip the test and store function inline so far. - let some decl ← getStage1Decl? declName | return none - pure decl.getArity - | _ => - match (← findLambda? f) with - | none => return none - | some localDecl => - unless (← shouldInline localDecl) do return none - pure (getLambdaArity localDecl.value) - if e.getAppNumArgs < arity then return none - return e.getAppNumArgs - arity + if let .const declName us ← findExpr f then + unless hasInlineAttribute (← getEnv) declName do return none + -- TODO: check whether function is recursive or not. + -- We can skip the test and store function inline so far. + let some decl ← getStage1Decl? declName | return none + let numArgs := e.getAppNumArgs + let arity := decl.getArity + if numArgs < arity then return none + /- + Recall that we use binder names to build `InlineStats`. + Thus, we use `ensureUniqueLetVarNames` to make sure there is no name collision. + -/ + let value ← ensureUniqueLetVarNames (decl.value.instantiateLevelParams decl.levelParams us) + return some { + arity, value + isLocal := false + } + else if let some localDecl ← findLambda? f then + unless (← shouldInlineLocal localDecl) do return none + let numArgs := e.getAppNumArgs + let arity := getLambdaArity localDecl.value + if numArgs < arity then return none + let value ← ensureUniqueLetVarNames localDecl.value + return some { + arity, value + isLocal := true + } + else + return none /-- If `e` if a free variable that expands to a valid LCNF terminal `let`-block expression `e'`, -return `e'`. -/ +return `e'`. +-/ def expandTrivialExpr (e : Expr) : SimpM Expr := do if e.isFVar then let e' ← findExpr e unless e'.isLambda do - if e != e' then markSimplified - return e' + if e != e' then + markSimplified + return e' return e -mutual +/-- +Given `value` of the form `let x_1 := v_1; ...; let x_n := v_n; e`, +return `let x_1; ...; let x_n := v_n; let y : type := e; body`. + +This methods assumes `type` and `value` do not have loose bound variables. + +Remark: `body` may have many loose bound variables, and the loose bound variables > 0 +must be lifted by `n`. +-/ +def mkFlatLet (y : Name) (type : Expr) (value : Expr) (body : Expr) (nonDep : Bool := false) : Expr := + go value 0 +where + go (value : Expr) (i : Nat) : Expr := + match value with + | .letE n t v b d => .letE n t v (go b (i+1)) d + | _ => .letE y type value (body.liftLooseBVars 1 i) nonDep -partial def visitLambda (e : Expr) : SimpM Expr := +/-- +Auxiliary function for projecting "type class dictionary access". +That is, we are trying to extract one of the type class instance elements. +Remark: We do not consider parent instances to be elements. +For example, suppose `e` is `_x_4.1`, and we have +``` +_x_2 : Monad (ReaderT Bool (ExceptT String Id)) := @ReaderT.Monad Bool (ExceptT String Id) _x_1 +_x_3 : Applicative (ReaderT Bool (ExceptT String Id)) := _x_2.1 +_x_4 : Functor (ReaderT Bool (ExceptT String Id)) := _x_3.1 +``` +Then, we will expand `_x_4.1` since it corresponds to the `Functor` `map` element, +and its type is not a type class, but is of the form +``` +{α β : Type u} → (α → β) → ... +``` +In the example above, the compiler should not expand `_x_3.1` or `_x_2.1` because they are +type class applications: `Functor` and `Applicative` respectively. +By eagerly expanding them, we may produce inefficient and bloated code. +For example, we may be using `_x_3.1` to invoke a function that expects a `Functor` instance. +By expanding `_x_3.1` we will be just expanding the code that creates this instance. +-/ +partial def inlineProjInst? (e : Expr) : OptionT SimpM Expr := do + let .proj _ _ s := e | failure + let sType ← inferType s + guard (← isClass? sType).isSome + let eType ← inferType e + guard (← isClass? eType).isNone + /- + We use `withNewScope` + `mkLetUsingScope` to filter the relevant let-declarations. + Recall that we are extracting only one of the type class elements. + -/ + let value ← withNewScope do mkLetUsingScope (← visitProj e) + markSimplified + internalize value +where + visitProj (e : Expr) : OptionT SimpM Expr := do + let .proj _ i s := e | unreachable! + let s ← visit s + if let some (ctorVal, ctorArgs) := s.constructorApp? (← getEnv) then + return ctorArgs[ctorVal.numParams + i]! + else + failure + + visit (e : Expr) : OptionT SimpM Expr := do + let e ← findExpr e + if e.isConstructorApp (← getEnv) then + return e + else if e.isProj then + /- We may have nested projections as we traverse parent classes. -/ + visit (← visitProj e) + else + let .const declName us := e.getAppFn | failure + let some decl ← getStage1Decl? declName | failure + guard <| decl.getArity == e.getAppNumArgs + let value := decl.value.instantiateLevelParams decl.levelParams us + let value := value.beta e.getAppArgs + /- + Here, we just go inside of the let-declaration block without trying to simplify it. + Reason: a type class instannce may have many elements, and it does not make sense to simplify + all of them when we are extracting only one of them. + -/ + let value ← Compiler.visitLet (m := SimpM) value fun _ value => return value + visit value + +def betaReduce (e : Expr) (args : Array Expr) : SimpM Expr := do + -- TODO: add necessary casts + internalize (e.beta args) + +/-- +Try "cases on cases" simplification. +If `casesFn args` is of the form +``` +casesOn _x.i + (... let _x.j₁ := ctorⱼ₁ ...; _jp.k _x.j₁) + ... + (... let _x.jₙ := ctorⱼₙ ...; _jp.k _x.jₙ) +``` +where `_jp.k` is a join point of the form +``` +let _jp.k := fun y => + casesOn y ... +``` +Then, inline `_jp.k`. The idea is to force the `casesOn` application in the join point to +reduce after the inlining step. +Example: consider the following declarations +``` +@[inline] def pred? (x : Nat) : Option Nat := + match x with + | 0 => none + | x+1 => some x + +def isZero (x : Nat) := + match pred? x with + | some _ => false + | none => true +``` +After inlining `pred?` in `isZero`, we have +``` +let _jp.1 := fun y : Option Nat => + casesOn y true (fun y => false) +casesOn x + (let _x.1 := none; _jp.1 _x.1) + (fun n => let _x.2 := some n; _jp.1 _x.2) +``` +and this simplification is applicable, producing +``` +casesOn x true (fun n => false) +``` +-/ +def simpCasesOnCases? (casesInfo : CasesInfo) (casesFn : Expr) (args : Array Expr) : OptionT SimpM Expr := do + let mut jpFirst? := none + for i in casesInfo.altsRange do + let alt := args[i]! + let jp ← isJpCtor? alt + if let some jpFirst := jpFirst? then + guard <| jp == jpFirst + else + let some localDecl ← findDecl? jp | failure + let .lam _ _ jpBody _ := localDecl.value | failure + guard (← isCasesApp? jpBody).isSome + jpFirst? := jp + let some jpFVarId := jpFirst? | failure + let some localDecl ← findDecl? jpFVarId | failure + let .lam _ _ jpBody _ := localDecl.value | failure + let mut args := args + for i in casesInfo.altsRange do + args := args.modify i (inlineJp · jpBody) + return mkAppN casesFn args +where + isJpCtor? (alt : Expr) : OptionT SimpM FVarId := do + match alt with + | .lam _ _ b _ => isJpCtor? b + | .letE _ _ v b _ => match b with + | .letE .. => isJpCtor? b + | .app (.fvar fvarId) (.bvar 0) => + let some localDecl ← findDecl? fvarId | failure + guard localDecl.isJp + guard <| v.isConstructorApp (← getEnv) + return fvarId + | _ => failure + | _ => failure + + inlineJp (alt : Expr) (jpBody : Expr) : Expr := + match alt with + | .lam n d b bi => .lam n d (inlineJp b jpBody) bi + | .letE n t v b nd => .letE n t v (inlineJp b jpBody) nd + | _ => jpBody + +mutual +/-- +Simplify the given lambda expression. +If `checkEmptyTypes := true`, then return `fun a_i : t_i => lcUnreachable` if +`t_i` is the `Empty` type. +-/ +partial def visitLambda (e : Expr) (checkEmptyTypes := false): SimpM Expr := withNewScope do - let (as, e) ← Compiler.visitLambda e - let e ← mkLetUsingScope (← visitLet e) + let (as, e) ← Compiler.visitLambdaCore e + if checkEmptyTypes then + for a in as do + if (← isEmptyType (← inferType a)) then + let e := e.instantiateRev as + let unreach ← mkLcUnreachable (← inferType e) + let r ← mkLambda as unreach + return r + let e ← mkLetUsingScope (← visitLet e as) mkLambda as e partial def visitCases (casesInfo : CasesInfo) (e : Expr) : SimpM Expr := do + let f := e.getAppFn let mut args := e.getAppArgs let major := args[casesInfo.discrsRange.stop - 1]! let major ← findExpr major @@ -179,28 +414,12 @@ partial def visitCases (casesInfo : CasesInfo) (e : Expr) : SimpM Expr := do assert! !alt.isLambda markSimplified visitLet alt + else if let some e ← simpCasesOnCases? casesInfo f args then + visitCases casesInfo e else for i in casesInfo.altsRange do - args ← args.modifyM i visitLambda - return mkAppN e.getAppFn args - -partial def inlineApp (e : Expr) (jp? : Option Expr := none) : SimpM Expr := do - let f := e.getAppFn - trace[Compiler.simp.inline] "inlining {e}" - let value ← match f with - | .const declName us => - let some decl ← getStage1Decl? declName | unreachable! - pure <| decl.value.instantiateLevelParams decl.levelParams us - | _ => - let some localDecl ← findLambda? f | unreachable! - pure localDecl.value - let args := e.getAppArgs - let value := value.beta args - let value ← attachOptJp value jp? - assert! !value.isLambda - markSimplified - withReader (fun ctx => { ctx with localInline := !f.isConst }) do - visitLet value + args ← args.modifyM i (visitLambda · (checkEmptyTypes := true)) + return mkAppN f args /-- If `e` is an application that can be inlined, inline it. @@ -210,36 +429,46 @@ that need to instantiated with `xs`. That is, if `k? = some k`, then `k.instanti is an expression without loose bound variables. -/ partial def inlineApp? (e : Expr) (xs : Array Expr) (k? : Option Expr) : SimpM (Option Expr) := do - let some numExtraArgs ← inlineCandidate? e | return none + let some info ← inlineCandidate? e | return none let args := e.getAppArgs - if k?.isNone && numExtraArgs == 0 then - -- Easy case, there is not continuation and `e` is not over applied - inlineApp e + let numArgs := args.size + trace[Compiler.simp.inline] "inlining {e}" + markSimplified + if k?.isNone && numArgs == info.arity then + /- Easy case, there is no continuation and `e` is not over applied -/ + visitLet (← betaReduce info.value args) + else if (← onlyOneExitPoint info.value) then + /- If `info.value` has only one exit point, we don't need to create a new auxiliary join point -/ + let mut value ← betaReduce info.value args[:info.arity] + if numArgs > info.arity then + let type ← inferType (mkAppN e.getAppFn args[:info.arity]) + value := mkFlatLet (← mkAuxLetDeclName) type value (mkAppN (.bvar 0) args[info.arity:]) + if let some k := k? then + let type ← inferType e + value := mkFlatLet (← mkAuxLetDeclName) type value k + visitLet value xs else /- There is a continuation `k` or `e` is over applied. - If `e` is over applied, the extra arguments act as continuation. - -/ - let toInline := mkAppN e.getAppFn args[:args.size - numExtraArgs] - /- - `toInline` is the application that is going to be inline - We create a new join point - ``` - let jp := fun y => - let x := y -- if `e` is over applied - k - ``` - Recall that `visitLet` incorporates the current continuation - to the new join point `jp`. + If `e` is over applied, the extra arguments act as a continuation. + + We create a new join point + ``` + let jp := fun y => + let x := y -- if `e` is over applied + k + ``` + Recall that `visitLet` incorporates the current continuation + to the new join point `jp`. -/ - let jpDomain ← inferType toInline + let jpDomain ← inferType (mkAppN e.getAppFn args[:info.arity]) let binderName ← mkFreshUserName `_y let jp ← withNewScope do let y ← mkLocalDecl binderName jpDomain - let body ← if numExtraArgs == 0 then + let body ← if numArgs == info.arity then visitLet k?.get! (xs.push y) else - let x ← mkAuxLetDecl (mkAppN y args[args.size - numExtraArgs:]) + let x ← mkAuxLetDecl (mkAppN y args[info.arity:]) if let some k := k? then visitLet k (xs.push x) else @@ -247,8 +476,13 @@ partial def inlineApp? (e : Expr) (xs : Array Expr) (k? : Option Expr) : SimpM ( let body ← mkLetUsingScope body mkLambda #[y] body let jp ← mkJpDeclIfNotSimple jp - /- Inline `toInline` and "go-to" `jp` with the result. -/ - inlineApp toInline jp + let value ← betaReduce info.value args[:info.arity] + let value ← attachJp value jp + visitLet value + +/-- Try to apply simple simplifications. -/ +partial def simpValue? (e : Expr) : SimpM (Option Expr) := + simpProj? e <|> simpAppApp? e <|> inlineProjInst? e /-- Let-declaration basic block visitor. `e` may contain loose bound variables that @@ -260,6 +494,12 @@ partial def visitLet (e : Expr) (xs : Array Expr := #[]): SimpM Expr := do let mut value := value.instantiateRev xs if value.isLambda then value ← visitLambda value + else if let some value' ← simpValue? value then + if value'.isLet then + let e := mkFlatLet binderName type value' body nonDep + let e ← visitLet e xs + return e + value := value' if value.isFVar then /- Eliminate `let _x_i := _x_j;` -/ markSimplified @@ -272,7 +512,9 @@ partial def visitLet (e : Expr) (xs : Array Expr := #[]): SimpM Expr := do visitLet body (xs.push x) | _ => let e := e.instantiateRev xs - if let some casesInfo ← isCasesApp? e then + if let some value ← simpValue? e then + visitLet value + else if let some casesInfo ← isCasesApp? e then visitCases casesInfo e else if let some e ← inlineApp? e #[] none then return e @@ -282,18 +524,20 @@ end end Simp -def Decl.simp? (decl : Decl) : CoreM (Option Decl) := do - let decl ← decl.ensureUniqueLetVarNames - let stats ← Simp.collectInlineStats decl.value - trace[Compiler.simp.inline.stats] "{decl.name}:{Format.nest 2 (format stats)}" - let (value, s) ← Simp.visitLambda decl.value |>.run { stats } |>.run { simplified := false } |>.run' {} - if s.simplified then +def Decl.simp? (decl : Decl) : Simp.SimpM (Option Decl) := do + let value ← Simp.internalize decl.value + trace[Compiler.simp.inline.occs] "{decl.name}:{Format.nest 2 (format (← get).occInfo)}" + trace[Compiler.simp.step] "{decl.name} :=\n{decl.value}" + let value ← Simp.visitLambda value + trace[Compiler.simp.step.new] "{decl.name} :=\n{value}" + trace[Compiler.simp.stat] "{decl.name}: {← getLCNFSize decl.value}" + if (← get).simplified then return some { decl with value } else return none partial def Decl.simp (decl : Decl) : CoreM Decl := do - if let some decl ← decl.simp? then + if let some decl ← decl.simp? |>.run {} |>.run' {} |>.run' {} then -- TODO: bound number of steps? decl.simp else @@ -301,7 +545,9 @@ partial def Decl.simp (decl : Decl) : CoreM Decl := do builtin_initialize registerTraceClass `Compiler.simp.inline + registerTraceClass `Compiler.simp.stat registerTraceClass `Compiler.simp.step - registerTraceClass `Compiler.simp.inline.stats + registerTraceClass `Compiler.simp.step.new + registerTraceClass `Compiler.simp.inline.occs -end Lean.Compiler \ No newline at end of file +end Lean.Compiler diff --git a/stage0/src/Lean/Compiler/Stage1.lean b/stage0/src/Lean/Compiler/Stage1.lean index 86b7e265ca80..327ca4ea3cf2 100644 --- a/stage0/src/Lean/Compiler/Stage1.lean +++ b/stage0/src/Lean/Compiler/Stage1.lean @@ -36,7 +36,10 @@ def getStage1Decl? (declName : Name) : CoreM (Option Decl) := do | some decl => return decl | none => let info ← getConstInfo declName - let decls ← compileStage1 info.all.toArray - return decls.find? fun decl => declName == decl.name + if info.hasValue then + let decls ← compileStage1 info.all.toArray + return decls.find? fun decl => declName == decl.name + else + return none end Lean.Compiler \ No newline at end of file diff --git a/stage0/src/Lean/Compiler/TerminalCases.lean b/stage0/src/Lean/Compiler/TerminalCases.lean index 539067e697a9..90859b6cfb96 100644 --- a/stage0/src/Lean/Compiler/TerminalCases.lean +++ b/stage0/src/Lean/Compiler/TerminalCases.lean @@ -21,8 +21,8 @@ private partial def visitCases (casesInfo : CasesInfo) (cases : Expr) : Compiler private partial def visitLambda (e : Expr) : CompilerM Expr := withNewScope do - let (as, e) ← Compiler.visitLambda e - let e ← mkLetUsingScope (← visitLet e #[]) + let (as, e) ← Compiler.visitLambdaCore e + let e ← mkLetUsingScope (← visitLet e as) mkLambda as e private partial def visitLet (e : Expr) (fvars : Array Expr) : CompilerM Expr := do diff --git a/stage0/src/Lean/Compiler/Util.lean b/stage0/src/Lean/Compiler/Util.lean index 71fa94e1070d..467cbf29fc8f 100644 --- a/stage0/src/Lean/Compiler/Util.lean +++ b/stage0/src/Lean/Compiler/Util.lean @@ -146,6 +146,30 @@ where for i in casesInfo.altsRange do go args[i]! +/-- +Return `true` if `getLCNFSIze e ≤ n` +-/ +partial def lcnfSizeLe (e : Expr) (n : Nat) : CoreM Bool := do + go e |>.run' 0 +where + inc : StateRefT Nat CoreM Bool := do + modify (·+1) + return (← get) <= n + + go (e : Expr) : StateRefT Nat CoreM Bool := do + match e with + | .lam _ _ b _ => go b + | .letE _ _ v b _ => inc <&&> go v <&&> go b + | _ => + unless (← inc) do + return false + if let some casesInfo ← isCasesApp? e then + let args := e.getAppArgs + for i in casesInfo.altsRange do + unless (← go args[i]!) do + return false + return true + def getLambdaArity (e : Expr) := match e with | .lam _ _ b _ => getLambdaArity b + 1 @@ -175,4 +199,33 @@ def isJump? [Monad m] [MonadLCtx m] (e : Expr) : m (Option FVarId) := do else return none +/-- +Return `true` if the LCNF expression has many exit points. +It assumes `cases` expressions only occur at the end of `let`-blocks. +That is, `terminalCases` has already been applied. +It also assumes that if contains a join point, then it has multiple +exit points. This is a reasonable assumption because the simplifier +inlines any join point that was used only once. +-/ +def manyExitPoints (e : Expr) : CoreM Bool := do + match e with + | .lam _ _ b _ => manyExitPoints b + | .letE n _ _ b _ => pure (isJpBinderName n) <||> manyExitPoints b + | e => return (← isCasesApp? e).isSome + +/-- +Return `true` if the LCNF expression has only one exit point. +-/ +def onlyOneExitPoint (e : Expr) : CoreM Bool := do + return !(← manyExitPoints e) + +/-- +Return `true` if `type` is an empty type. + +Remark: this is an approximate test that only checks +whether `type == Empty`. It is good enough (and fast) for our purposes. +-/ +def isEmptyType (type : Expr) : CoreM Bool := + return type.isConstOf ``Empty + end Lean.Compiler diff --git a/stage0/src/Lean/Elab/Quotation.lean b/stage0/src/Lean/Elab/Quotation.lean index 0d1603ff0272..c240193be7a3 100644 --- a/stage0/src/Lean/Elab/Quotation.lean +++ b/stage0/src/Lean/Elab/Quotation.lean @@ -99,6 +99,11 @@ def tryAddSyntaxNodeKindInfo (stx : Syntax) (k : SyntaxNodeKind) : TermElabM Uni if (← getEnv).contains k then addTermInfo' stx (← mkConstWithFreshMVarLevels k) +instance : Quote Syntax.Preresolved where + quote + | .namespace ns => Unhygienic.run ``(Syntax.Preresolved.namespace $(quote ns)) + | .decl n fs => Unhygienic.run ``(Syntax.Preresolved.decl $(quote n) $(quote fs)) + /-- Elaborate the content of a syntax quotation term -/ private partial def quoteSyntax : Syntax → TermElabM Term | Syntax.ident _ rawVal val preresolved => do @@ -109,8 +114,9 @@ private partial def quoteSyntax : Syntax → TermElabM Term let r ← resolveGlobalName val -- extension of the paper algorithm: also store unique section variable names as top-level scopes -- so they can be captured and used inside the section, but not outside - let r' := resolveSectionVariable (← read).sectionVars val - let preresolved := r ++ r' ++ preresolved + let r := r ++ resolveSectionVariable (← read).sectionVars val + let preresolved := r ++ preresolved + let preresolved := preresolved.map fun (n, fs) => Syntax.Preresolved.decl n fs let val := quote val -- `scp` is bound in stxQuot.expand `(Syntax.ident info $(quote rawVal) (addMacroScope mainModule $val scp) $(quote preresolved)) diff --git a/stage0/src/Lean/Meta/InferType.lean b/stage0/src/Lean/Meta/InferType.lean index 83c401fafe99..6249806b1244 100644 --- a/stage0/src/Lean/Meta/InferType.lean +++ b/stage0/src/Lean/Meta/InferType.lean @@ -151,12 +151,6 @@ private def inferLambdaType (e : Expr) : MetaM Expr := let type ← inferType e mkForallFVars xs type -@[inline] private def withLocalDecl' {α} (name : Name) (bi : BinderInfo) (type : Expr) (x : Expr → MetaM α) : MetaM α := - savingCache do - let fvarId ← mkFreshFVarId - withReader (fun ctx => { ctx with lctx := ctx.lctx.mkLocalDecl fvarId name type bi }) do - x (mkFVar fvarId) - def throwUnknownMVar {α} (mvarId : MVarId) : MetaM α := throwError "unknown metavariable '?{mvarId.name}'" @@ -377,16 +371,35 @@ def isType (e : Expr) : MetaM Bool := do | .sort .. => return true | _ => return false + +@[inline] private def withLocalDecl' {α} (name : Name) (bi : BinderInfo) (type : Expr) (x : Expr → MetaM α) : MetaM α := do + let fvarId ← mkFreshFVarId + withReader (fun ctx => { ctx with lctx := ctx.lctx.mkLocalDecl fvarId name type bi }) do + x (mkFVar fvarId) + +def isTypeFormerTypeQuick : Expr → Bool + | .forallE _ _ b _ => isTypeFormerTypeQuick b + | .sort _ => true + | _ => false + /-- Return true iff `type` is `Sort _` or `As → Sort _`. -/ partial def isTypeFormerType (type : Expr) : MetaM Bool := do - let type ← whnfD type - match type with - | .sort .. => return true - | .forallE n d b c => - withLocalDecl' n c d fun fvar => isTypeFormerType (b.instantiate1 fvar) - | _ => return false + match isTypeFormerTypeQuick type with + | true => return true + | false => savingCache <| go type #[] +where + go (type : Expr) (xs : Array Expr) : MetaM Bool := do + match type with + | .sort .. => return true + | .forallE n d b c => withLocalDecl' n c (d.instantiateRev xs) fun x => go b (xs.push x) + | _ => + let type ← whnfD (type.instantiateRev xs) + match type with + | .sort .. => return true + | .forallE .. => go type #[] + | _ => return false /-- Return true iff `e : Sort _` or `e : (forall As, Sort _)`. diff --git a/stage0/src/Std/Data/HashSet.lean b/stage0/src/Std/Data/HashSet.lean index aaf7926749d8..d63e7876861d 100644 --- a/stage0/src/Std/Data/HashSet.lean +++ b/stage0/src/Std/Data/HashSet.lean @@ -46,6 +46,12 @@ def mkIdx {n : Nat} (h : n > 0) (u : USize) : { u : USize // u.toNat < n } := @[inline] def fold {δ : Type w} (f : δ → α → δ) (d : δ) (m : HashSetImp α) : δ := foldBuckets m.buckets d f +@[inline] def forBucketsM {m : Type w → Type w} [Monad m] (data : HashSetBucket α) (f : α → m PUnit) : m PUnit := + data.val.forM fun as=> as.forM f + +@[inline] def forM {m : Type w → Type w} [Monad m] (f : α → m PUnit) (h : HashSetImp α) : m PUnit := + forBucketsM h.buckets f + def find? [BEq α] [Hashable α] (m : HashSetImp α) (a : α) : Option α := match m with | ⟨_, buckets⟩ => @@ -149,6 +155,13 @@ variable {α : Type u} {_ : BEq α} {_ : Hashable α} match m with | ⟨ m, _ ⟩ => m.fold f init +@[inline] def forM {m : Type w → Type w} [Monad m] (h : HashSet α) (f : α → m PUnit) : m PUnit := + match h with + | ⟨h, _⟩ => h.forM f + +instance : ForM m (HashSet α) α where + forM := HashSet.forM + @[inline] def size (m : HashSet α) : Nat := match m with | ⟨ {size := sz, ..}, _ ⟩ => sz diff --git a/stage0/stdlib/Lean/Compiler/CSE.c b/stage0/stdlib/Lean/Compiler/CSE.c index eb5402836b3d..b574353edcea 100644 --- a/stage0/stdlib/Lean/Compiler/CSE.c +++ b/stage0/stdlib/Lean/Compiler/CSE.c @@ -16,30 +16,24 @@ extern "C" { lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Compiler_CSE_State_map___default___closed__1; lean_object* l_Lean_Compiler_mkLetUsingScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSE_State_map___default___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Compiler_Decl_mapValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Compiler_withNewScopeImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_CSE_visitLet___closed__1; extern lean_object* l_Lean_levelZero; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSE_visitLambda___closed__1; +lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLambda___lambda__1(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_Std_Range_forIn_loop___at_Lean_Compiler_CSE_visitCases___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_st_ref_take(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkLetDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSE_visitCases___closed__1; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); @@ -50,16 +44,16 @@ lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_cse___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_Decl_cse___closed__1; uint8_t l_Lean_Expr_isLambda(lean_object*); -lean_object* l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_visitLambdaCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSE_State_map___default___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t l_Lean_Compiler_isJpBinderName(lean_object*); lean_object* l_Lean_Compiler_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet_go___lambda__2(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_Compiler_Decl_cse(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); +lean_object* l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_State_map___default; LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet_go___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*); @@ -67,7 +61,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet_go(lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_CSE_visitLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLet_go___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_Std_Range_forIn_loop___at_Lean_Compiler_CSE_visitCases___spec__1(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_lt(lean_object*, lean_object*); @@ -135,7 +128,7 @@ x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); lean_inc(x_6); -x_18 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_16, x_6); +x_18 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(x_16, x_6); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; @@ -165,7 +158,7 @@ x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); lean_inc(x_21); -x_29 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__1(x_27, x_6, x_21); +x_29 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(x_27, x_6, x_21); x_30 = lean_st_ref_set(x_8, x_29, x_28); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); @@ -573,218 +566,480 @@ return x_32; } } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___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) { +static lean_object* _init_l_Lean_Compiler_CSE_visitLambda___closed__1() { _start: { -lean_object* x_8; -lean_inc(x_6); +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_Lean_Compiler_CSE_visitLambda(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; lean_object* x_12; lean_object* x_13; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_3, 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_40 = lean_st_ref_get(x_5, x_11); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_st_ref_take(x_3, x_41); +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_is_exclusive(x_43); +if (x_45 == 0) +{ +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_46 = lean_ctor_get(x_43, 1); +lean_dec(x_46); +x_47 = l_Lean_Compiler_CSE_visitLambda___closed__1; +lean_ctor_set(x_43, 1, x_47); +x_48 = lean_st_ref_set(x_3, x_43, x_44); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_Lean_Compiler_visitLambdaCore(x_1, x_3, x_4, x_5, x_49); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_8 = lean_apply_5(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -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_apply_6(x_2, x_9, x_3, x_4, x_5, x_6, x_10); -return x_11; -} -else -{ -uint8_t x_12; -lean_dec(x_6); -lean_dec(x_5); +lean_inc(x_53); +x_55 = l_Lean_Compiler_CSE_visitLet(x_54, x_53, x_2, x_3, x_4, x_5, x_52); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +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); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_58 = l_Lean_Compiler_mkLetUsingScope(x_56, x_3, x_4, x_5, x_57); +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; 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; +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 = l_Lean_Compiler_mkLambda(x_53, x_59, x_3, x_4, x_5, x_60); lean_dec(x_4); +lean_dec(x_53); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_st_ref_get(x_5, x_63); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = lean_st_ref_get(x_3, x_65); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_10, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_10, 1); +lean_inc(x_70); +lean_dec(x_10); +x_71 = !lean_is_exclusive(x_67); +if (x_71 == 0) +{ +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_67, 1); +lean_dec(x_72); +x_73 = lean_ctor_get(x_67, 0); +lean_dec(x_73); +lean_ctor_set(x_67, 1, x_70); +lean_ctor_set(x_67, 0, x_69); +x_74 = lean_st_ref_get(x_5, x_68); +lean_dec(x_5); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = lean_st_ref_set(x_3, x_67, x_75); lean_dec(x_3); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_8); -if (x_12 == 0) +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) { -return x_8; +lean_object* x_78; +x_78 = lean_ctor_get(x_76, 0); +lean_dec(x_78); +lean_ctor_set(x_76, 0, x_62); +return x_76; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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 = 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_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_76, 1); +lean_inc(x_79); +lean_dec(x_76); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_62); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } +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; +x_81 = lean_ctor_get(x_67, 2); +lean_inc(x_81); +lean_dec(x_67); +x_82 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_82, 0, x_69); +lean_ctor_set(x_82, 1, x_70); +lean_ctor_set(x_82, 2, x_81); +x_83 = lean_st_ref_get(x_5, x_68); +lean_dec(x_5); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = lean_st_ref_set(x_3, x_82, x_84); +lean_dec(x_3); +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_87 = x_85; +} else { + lean_dec_ref(x_85); + x_87 = lean_box(0); +} +if (lean_is_scalar(x_87)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_87; +} +lean_ctor_set(x_88, 0, x_62); +lean_ctor_set(x_88, 1, x_86); +return x_88; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___spec__1(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___spec__1___rarg), 7, 0); -return x_3; +lean_object* x_89; lean_object* x_90; +lean_dec(x_53); +lean_dec(x_4); +x_89 = lean_ctor_get(x_58, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_58, 1); +lean_inc(x_90); +lean_dec(x_58); +x_12 = x_89; +x_13 = x_90; +goto block_39; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_CSE_visitLambda___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) { -_start: +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_apply_1(x_1, x_2); -x_8 = l_Lean_Compiler_withNewScopeImp___rarg(x_7, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_91; lean_object* x_92; +lean_dec(x_53); +lean_dec(x_4); +x_91 = lean_ctor_get(x_55, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_55, 1); +lean_inc(x_92); +lean_dec(x_55); +x_12 = x_91; +x_13 = x_92; +goto block_39; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLambda___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: +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); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_dec(x_1); +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_93 = lean_ctor_get(x_43, 0); +x_94 = lean_ctor_get(x_43, 2); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_43); +x_95 = l_Lean_Compiler_CSE_visitLambda___closed__1; +x_96 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_96, 0, x_93); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = lean_st_ref_set(x_3, x_96, x_44); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_99 = l_Lean_Compiler_visitLambdaCore(x_1, x_3, x_4, x_5, x_98); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_102 = lean_ctor_get(x_100, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_100, 1); +lean_inc(x_103); +lean_dec(x_100); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_9 = l_Lean_Compiler_CSE_visitLet(x_8, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -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_102); +x_104 = l_Lean_Compiler_CSE_visitLet(x_103, x_102, x_2, x_3, x_4, x_5, x_101); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_12 = l_Lean_Compiler_mkLetUsingScope(x_10, x_3, x_4, x_5, 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 = l_Lean_Compiler_mkLambda(x_7, x_13, x_3, x_4, x_5, x_14); -lean_dec(x_5); +x_107 = l_Lean_Compiler_mkLetUsingScope(x_105, x_3, x_4, x_5, x_106); +if (lean_obj_tag(x_107) == 0) +{ +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; +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +lean_dec(x_107); +x_110 = l_Lean_Compiler_mkLambda(x_102, x_108, x_3, x_4, x_5, x_109); lean_dec(x_4); +lean_dec(x_102); +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_st_ref_get(x_5, x_112); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +lean_dec(x_113); +x_115 = lean_st_ref_get(x_3, x_114); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_ctor_get(x_10, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_10, 1); +lean_inc(x_119); +lean_dec(x_10); +x_120 = lean_ctor_get(x_116, 2); +lean_inc(x_120); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + lean_ctor_release(x_116, 2); + 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, 3, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_118); +lean_ctor_set(x_122, 1, x_119); +lean_ctor_set(x_122, 2, x_120); +x_123 = lean_st_ref_get(x_5, x_117); +lean_dec(x_5); +x_124 = lean_ctor_get(x_123, 1); +lean_inc(x_124); +lean_dec(x_123); +x_125 = lean_st_ref_set(x_3, x_122, x_124); lean_dec(x_3); -lean_dec(x_7); -return x_15; +x_126 = lean_ctor_get(x_125, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_127 = x_125; +} else { + lean_dec_ref(x_125); + x_127 = lean_box(0); +} +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(0, 2, 0); +} else { + x_128 = x_127; +} +lean_ctor_set(x_128, 0, x_111); +lean_ctor_set(x_128, 1, x_126); +return x_128; } else { -uint8_t x_16; -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_129; lean_object* x_130; +lean_dec(x_102); lean_dec(x_4); -lean_dec(x_3); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -return x_12; +x_129 = lean_ctor_get(x_107, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_107, 1); +lean_inc(x_130); +lean_dec(x_107); +x_12 = x_129; +x_13 = x_130; +goto block_39; +} } 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_object* x_131; lean_object* x_132; +lean_dec(x_102); +lean_dec(x_4); +x_131 = lean_ctor_get(x_104, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_104, 1); +lean_inc(x_132); +lean_dec(x_104); +x_12 = x_131; +x_13 = x_132; +goto block_39; } } -else +block_39: { -uint8_t x_20; -lean_dec(x_7); +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_14 = lean_st_ref_get(x_5, x_13); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_3, x_15); +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_10, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_10, 1); +lean_inc(x_20); +lean_dec(x_10); +x_21 = !lean_is_exclusive(x_17); +if (x_21 == 0) +{ +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_ctor_get(x_17, 1); +lean_dec(x_22); +x_23 = lean_ctor_get(x_17, 0); +lean_dec(x_23); +lean_ctor_set(x_17, 1, x_20); +lean_ctor_set(x_17, 0, x_19); +x_24 = lean_st_ref_get(x_5, x_18); lean_dec(x_5); -lean_dec(x_4); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_st_ref_set(x_3, x_17, x_25); lean_dec(x_3); -x_20 = !lean_is_exclusive(x_9); -if (x_20 == 0) +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -return x_9; +lean_object* x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +lean_ctor_set_tag(x_26, 1); +lean_ctor_set(x_26, 0, x_12); +return x_26; } 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; -} -} -} -} -static lean_object* _init_l_Lean_Compiler_CSE_visitLambda___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_CSE_visitLambda___lambda__1), 6, 0); -return x_1; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_12); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_CSE_visitLambda(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: +else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_7, 0, x_1); -x_8 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Compiler_CSE_visitLambda___closed__1; -x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_CSE_visitLambda___spec__1___rarg), 7, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_CSE_visitLambda___spec__2(x_10, x_2, x_3, x_4, x_5, x_6); -return x_11; +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_31 = lean_ctor_get(x_17, 2); +lean_inc(x_31); +lean_dec(x_17); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_19); +lean_ctor_set(x_32, 1, x_20); +lean_ctor_set(x_32, 2, x_31); +x_33 = lean_st_ref_get(x_5, x_18); +lean_dec(x_5); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = lean_st_ref_set(x_3, x_32, x_34); +lean_dec(x_3); +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_35); + x_37 = lean_box(0); +} +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(1, 2, 0); +} else { + x_38 = x_37; + lean_ctor_set_tag(x_38, 1); +} +lean_ctor_set(x_38, 0, x_12); +lean_ctor_set(x_38, 1, x_36); +return x_38; } } -static lean_object* _init_l_Lean_Compiler_CSE_visitLet___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_Lean_Compiler_CSE_visitLet(lean_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_Compiler_CSE_visitLet(lean_object* x_1, 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; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_st_ref_get(x_5, x_6); -x_8 = lean_ctor_get(x_7, 1); -lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_st_ref_get(x_2, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +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_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Compiler_CSE_visitLet___closed__1; -lean_inc(x_5); -lean_inc(x_2); -x_13 = l_Lean_Compiler_CSE_visitLet_go(x_1, x_12, x_2, x_3, x_4, x_5, x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_6); +lean_inc(x_3); +x_13 = l_Lean_Compiler_CSE_visitLet_go(x_1, x_2, x_3, x_4, x_5, x_6, x_12); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -793,13 +1048,13 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_st_ref_get(x_5, x_15); -lean_dec(x_5); +x_16 = lean_st_ref_get(x_6, x_15); +lean_dec(x_6); x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = lean_st_ref_set(x_2, x_10, x_17); -lean_dec(x_2); +x_18 = lean_st_ref_set(x_3, x_11, x_17); +lean_dec(x_3); x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) { @@ -829,13 +1084,13 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_13, 1); lean_inc(x_24); lean_dec(x_13); -x_25 = lean_st_ref_get(x_5, x_24); -lean_dec(x_5); +x_25 = lean_st_ref_get(x_6, x_24); +lean_dec(x_6); x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); lean_dec(x_25); -x_27 = lean_st_ref_set(x_2, x_10, x_26); -lean_dec(x_2); +x_27 = lean_st_ref_set(x_3, x_11, x_26); +lean_dec(x_3); x_28 = !lean_is_exclusive(x_27); if (x_28 == 0) { @@ -1006,8 +1261,6 @@ l_Lean_Compiler_CSE_visitCases___closed__1 = _init_l_Lean_Compiler_CSE_visitCase lean_mark_persistent(l_Lean_Compiler_CSE_visitCases___closed__1); l_Lean_Compiler_CSE_visitLambda___closed__1 = _init_l_Lean_Compiler_CSE_visitLambda___closed__1(); lean_mark_persistent(l_Lean_Compiler_CSE_visitLambda___closed__1); -l_Lean_Compiler_CSE_visitLet___closed__1 = _init_l_Lean_Compiler_CSE_visitLet___closed__1(); -lean_mark_persistent(l_Lean_Compiler_CSE_visitLet___closed__1); l_Lean_Compiler_Decl_cse___closed__1 = _init_l_Lean_Compiler_Decl_cse___closed__1(); lean_mark_persistent(l_Lean_Compiler_Decl_cse___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Compiler/Check.c b/stage0/stdlib/Lean/Compiler/Check.c index 58d0363cc462..0113a5d8df59 100644 --- a/stage0/stdlib/Lean/Compiler/Check.c +++ b/stage0/stdlib/Lean/Compiler/Check.c @@ -13,10 +13,9 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__2; static lean_object* l_Lean_Compiler_check_checkBlock___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___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_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__10(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Compiler_InferType_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -24,7 +23,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_check_checkBlock___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -44,7 +43,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTeles lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkValue(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkValue(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_check_checkBlock___closed__7; static lean_object* l_Lean_Compiler_check_checkBlock___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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*); @@ -52,7 +51,7 @@ uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__6(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,10 +63,8 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkCases___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_check_checkApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg___closed__3; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___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_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg___closed__5; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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_Compiler_check_checkBlock___lambda__3___closed__8; static lean_object* l_Lean_Compiler_check_checkBlock___lambda__3___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__1; @@ -76,17 +73,16 @@ static lean_object* l_Lean_Compiler_check_checkValue___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__2(lean_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg___closed__1; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___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_Compiler_check_checkCases(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkCases(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_Compiler_check_checkBlock___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___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_Lean_Compiler_check_checkValue___closed__1; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__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___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___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* l_Std_RBNode_setBlack___rarg(lean_object*); @@ -97,7 +93,7 @@ lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev_range(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_check_checkBlock___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__1(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_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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_object*); @@ -112,6 +108,7 @@ lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__2___closed__6; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_Check_Config_checkJpScope___default; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); uint8_t l_Lean_Compiler_isTypeFormerType(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__3___closed__1; @@ -142,7 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_withFreshJps(lean_object*); lean_object* lean_local_ctx_mk_local_decl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_check___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkLambda(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_lambdaBoundedTelescope(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_lambdaBoundedTelescope_go(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_withFreshJps___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -159,13 +156,12 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___s static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__1(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_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2___closed__1; static lean_object* l_Lean_Compiler_check_checkApp___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___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_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___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_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___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_Compiler_check_checkBlock(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2___closed__2; uint8_t l_Lean_LocalDecl_isJp(lean_object*); static lean_object* l_Lean_Compiler_lambdaBoundedTelescope___rarg___closed__1; @@ -175,7 +171,6 @@ uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_check_checkApp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_check_checkApp___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_check_checkBlock___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___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*); @@ -192,7 +187,7 @@ static lean_object* l_Lean_Compiler_check_checkBlock___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg___closed__6; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__4(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_Compiler_check_checkBlock___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); @@ -209,18 +204,25 @@ LEAN_EXPORT lean_object* l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlo static lean_object* l_Lean_Compiler_check_checkApp___closed__2; lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___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_Compiler_check_checkBlock___lambda__2(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_Compiler_check_checkBlock___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_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___lambda__4___closed__1; static lean_object* l_Lean_Compiler_lambdaBoundedTelescope_go___rarg___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__1___closed__2; static lean_object* l_Lean_Compiler_check_checkBlock___lambda__3___closed__5; -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkLambda___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_Compiler_lambdaBoundedTelescope_go___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_lambdaBoundedTelescope_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__1___closed__4; +static uint8_t _init_l_Lean_Compiler_Check_Config_checkJpScope___default() { +_start: +{ +uint8_t x_1; +x_1 = 1; +return x_1; +} +} static uint8_t _init_l_Lean_Compiler_Check_Config_terminalCasesOnly___default() { _start: { @@ -1571,7 +1573,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_lambdaBoundedTelescope___rarg), return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkLambda(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_EXPORT lean_object* l_Lean_Compiler_check_checkLambda(lean_object* x_1, lean_object* x_2, 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_2) == 6) @@ -4465,7 +4467,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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; lean_object* x_12; @@ -4478,7 +4480,7 @@ x_12 = l_Lean_Compiler_check_checkValue(x_3, x_10, x_11, x_5, x_6, x_7, x_8, x_9 return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__2(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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -4669,7 +4671,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__3(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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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) { _start: { lean_object* x_13; lean_object* x_14; @@ -4695,6 +4697,7 @@ 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; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_dec(x_6); +lean_dec(x_5); lean_dec(x_2); x_18 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_18, 0, x_4); @@ -4763,6 +4766,7 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); x_38 = !lean_is_exclusive(x_14); @@ -4786,7 +4790,7 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___lambda__4(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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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) { _start: { uint8_t x_13; @@ -4805,6 +4809,7 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); +lean_inc(x_5); x_16 = l_Lean_Compiler_check_checkValue(x_5, x_3, x_15, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_16) == 0) { @@ -4825,6 +4830,7 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -4851,112 +4857,115 @@ return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = l_Lean_Compiler_lambdaBoundedTelescope___rarg___closed__1; -x_25 = lean_box(x_5); lean_inc(x_3); -x_26 = lean_alloc_closure((void*)(l_Lean_Compiler_check_checkLambda___boxed), 8, 3); -lean_closure_set(x_26, 0, x_25); -lean_closure_set(x_26, 1, x_3); -lean_closure_set(x_26, 2, x_24); +lean_inc(x_5); +x_25 = lean_alloc_closure((void*)(l_Lean_Compiler_check_checkLambda), 8, 3); +lean_closure_set(x_25, 0, x_5); +lean_closure_set(x_25, 1, x_3); +lean_closure_set(x_25, 2, x_24); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_27 = l_Lean_Compiler_withFreshJps___rarg(x_26, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_27) == 0) +x_26 = l_Lean_Compiler_withFreshJps___rarg(x_25, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 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); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Compiler_check_checkBlock___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_28, x_8, x_9, x_10, x_11, x_29); -return x_30; +lean_dec(x_26); +x_29 = l_Lean_Compiler_check_checkBlock___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9, x_10, x_11, x_28); +return x_29; } else { -uint8_t x_31; +uint8_t x_30; 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_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_27); -if (x_31 == 0) +x_30 = !lean_is_exclusive(x_26); +if (x_30 == 0) { -return x_27; +return x_26; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_27, 0); -x_33 = lean_ctor_get(x_27, 1); -lean_inc(x_33); +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_dec(x_27); -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_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; } } } } else { -lean_object* x_35; lean_object* x_36; -x_35 = l_Lean_Compiler_lambdaBoundedTelescope___rarg___closed__1; +lean_object* x_34; lean_object* x_35; +x_34 = l_Lean_Compiler_lambdaBoundedTelescope___rarg___closed__1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_36 = l_Lean_Compiler_check_checkLambda(x_5, x_3, x_35, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_5); +x_35 = l_Lean_Compiler_check_checkLambda(x_5, x_3, x_34, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_36, 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); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_Compiler_check_checkBlock___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_37, x_8, x_9, x_10, x_11, x_38); -return x_39; +lean_dec(x_35); +x_38 = l_Lean_Compiler_check_checkBlock___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_36, x_8, x_9, x_10, x_11, x_37); +return x_38; } else { -uint8_t x_40; +uint8_t x_39; 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_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_40 = !lean_is_exclusive(x_36); -if (x_40 == 0) +x_39 = !lean_is_exclusive(x_35); +if (x_39 == 0) { -return x_36; +return x_35; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_36, 0); -x_42 = lean_ctor_get(x_36, 1); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_35, 0); +x_41 = lean_ctor_get(x_35, 1); lean_inc(x_41); -lean_dec(x_36); -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_inc(x_40); +lean_dec(x_35); +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; } } } @@ -5048,13 +5057,14 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock(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_EXPORT lean_object* l_Lean_Compiler_check_checkBlock(lean_object* x_1, lean_object* x_2, 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: { switch (lean_obj_tag(x_2)) { case 6: { 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_9 = lean_expr_instantiate_rev(x_2, x_3); lean_dec(x_3); lean_dec(x_2); @@ -5103,6 +5113,7 @@ lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); lean_dec(x_3); +lean_dec(x_1); x_25 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_25, 0, x_16); x_26 = l_Lean_Compiler_check_checkBlock___closed__8; @@ -5156,51 +5167,68 @@ return x_39; } 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_ctor_get(x_36, 0); +uint8_t x_40; +x_40 = lean_ctor_get_uint8(x_1, 0); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_36); +x_41 = lean_ctor_get(x_35, 1); lean_inc(x_41); +lean_dec(x_35); +x_42 = lean_box(0); +x_43 = l_Lean_Compiler_check_checkBlock___lambda__1(x_2, x_3, x_1, x_42, x_4, x_5, x_6, x_7, x_41); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_35, 1); +lean_inc(x_44); +lean_dec(x_35); +x_45 = lean_ctor_get(x_36, 0); +lean_inc(x_45); lean_dec(x_36); -x_42 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_4, x_41); -lean_dec(x_41); -if (lean_obj_tag(x_42) == 0) +x_46 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_4, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_dec(x_3); lean_dec(x_2); -x_43 = l_Lean_Compiler_check_checkBlock___closed__2; -x_44 = l_Lean_throwError___at_Lean_Compiler_check_checkBlock___spec__4(x_43, x_4, x_5, x_6, x_7, x_40); +lean_dec(x_1); +x_47 = l_Lean_Compiler_check_checkBlock___closed__2; +x_48 = l_Lean_throwError___at_Lean_Compiler_check_checkBlock___spec__4(x_47, x_4, x_5, x_6, x_7, x_44); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -return x_44; +return x_48; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_44, 0); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_44); -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_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_48); +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 { -lean_object* x_49; lean_object* x_50; -lean_dec(x_42); -x_49 = lean_box(0); -x_50 = l_Lean_Compiler_check_checkBlock___lambda__1(x_2, x_3, x_1, x_49, x_4, x_5, x_6, x_7, x_40); -return x_50; +lean_object* x_53; lean_object* x_54; +lean_dec(x_46); +x_53 = lean_box(0); +x_54 = l_Lean_Compiler_check_checkBlock___lambda__1(x_2, x_3, x_1, x_53, x_4, x_5, x_6, x_7, x_44); +return x_54; +} } } } @@ -5224,7 +5252,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkValue(uint8_t 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_Compiler_check_checkValue(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: { switch (lean_obj_tag(x_2)) { @@ -5236,6 +5264,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); 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); @@ -5250,6 +5279,7 @@ lean_dec(x_6); lean_dec(x_5); 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); @@ -5270,6 +5300,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +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); @@ -5279,6 +5310,7 @@ return x_15; case 10: { lean_object* x_16; +lean_dec(x_1); x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); if (lean_obj_tag(x_16) == 1) @@ -5319,6 +5351,7 @@ return x_24; case 11: { lean_object* x_25; +lean_dec(x_1); x_25 = lean_ctor_get(x_2, 2); lean_inc(x_25); if (lean_obj_tag(x_25) == 1) @@ -5359,6 +5392,7 @@ return x_33; default: { 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_dec(x_1); x_34 = l_Lean_indentExpr(x_2); x_35 = l_Lean_Compiler_check_checkValue___closed__2; x_36 = lean_alloc_ctor(10, 2, 0); @@ -6249,7 +6283,7 @@ return x_53; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__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_EXPORT lean_object* l_Lean_Compiler_check_checkApp___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; @@ -6275,7 +6309,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__2(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_EXPORT lean_object* l_Lean_Compiler_check_checkApp___lambda__2(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) { _start: { lean_object* x_12; @@ -6292,6 +6326,7 @@ lean_inc(x_13); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; +lean_dec(x_4); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); @@ -6442,104 +6477,108 @@ else lean_dec(x_2); if (x_5 == 0) { -if (x_4 == 0) +uint8_t x_47; +x_47 = lean_ctor_get_uint8(x_4, 1); +if (x_47 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_dec(x_1); -x_47 = lean_ctor_get(x_12, 1); -lean_inc(x_47); -lean_dec(x_12); -x_48 = lean_ctor_get(x_13, 0); +x_48 = lean_ctor_get(x_12, 1); lean_inc(x_48); +lean_dec(x_12); +x_49 = lean_ctor_get(x_13, 0); +lean_inc(x_49); lean_dec(x_13); -x_49 = l_Lean_Compiler_check_checkCases(x_4, x_48, x_3, x_7, x_8, x_9, x_10, x_47); +x_50 = l_Lean_Compiler_check_checkCases(x_4, x_49, x_3, x_7, x_8, x_9, x_10, x_48); lean_dec(x_3); -return x_49; +return x_50; } 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_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_13); +lean_dec(x_4); lean_dec(x_3); -x_50 = lean_ctor_get(x_12, 1); -lean_inc(x_50); +x_51 = lean_ctor_get(x_12, 1); +lean_inc(x_51); lean_dec(x_12); -x_51 = l_Lean_indentExpr(x_1); -x_52 = l_Lean_Compiler_check_checkApp___lambda__2___closed__2; -x_53 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_51); -x_54 = l_Lean_Compiler_check_checkBlock___lambda__3___closed__8; -x_55 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_throwError___at_Lean_Compiler_check_checkBlock___spec__4(x_55, x_7, x_8, x_9, x_10, x_50); +x_52 = l_Lean_indentExpr(x_1); +x_53 = l_Lean_Compiler_check_checkApp___lambda__2___closed__2; +x_54 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l_Lean_Compiler_check_checkBlock___lambda__3___closed__8; +x_56 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_throwError___at_Lean_Compiler_check_checkBlock___spec__4(x_56, x_7, x_8, x_9, x_10, x_51); lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) { -return x_56; +return x_57; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_56, 0); -x_59 = lean_ctor_get(x_56, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_57, 0); +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_56); -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_dec(x_57); +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 { -lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec(x_1); -x_61 = lean_ctor_get(x_12, 1); -lean_inc(x_61); -lean_dec(x_12); -x_62 = lean_ctor_get(x_13, 0); +x_62 = lean_ctor_get(x_12, 1); lean_inc(x_62); +lean_dec(x_12); +x_63 = lean_ctor_get(x_13, 0); +lean_inc(x_63); lean_dec(x_13); -x_63 = l_Lean_Compiler_check_checkCases(x_4, x_62, x_3, x_7, x_8, x_9, x_10, x_61); +x_64 = l_Lean_Compiler_check_checkCases(x_4, x_63, x_3, x_7, x_8, x_9, x_10, x_62); lean_dec(x_3); -return x_63; +return x_64; } } } else { -uint8_t x_64; +uint8_t x_65; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_12); -if (x_64 == 0) +x_65 = !lean_is_exclusive(x_12); +if (x_65 == 0) { return x_12; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_12, 0); -x_66 = lean_ctor_get(x_12, 1); +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_12, 0); +x_67 = lean_ctor_get(x_12, 1); +lean_inc(x_67); lean_inc(x_66); -lean_inc(x_65); lean_dec(x_12); -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; +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; } } } @@ -6570,7 +6609,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp(uint8_t 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_Compiler_check_checkApp(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; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; @@ -6595,6 +6634,7 @@ 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; uint8_t x_25; lean_dec(x_16); lean_dec(x_9); +lean_dec(x_1); x_19 = l_Lean_indentExpr(x_2); x_20 = l_Lean_Compiler_check_checkApp___closed__3; x_21 = lean_alloc_ctor(10, 2, 0); @@ -6908,7 +6948,7 @@ return x_65; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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) { _start: { lean_object* x_9; lean_object* x_10; @@ -6935,7 +6975,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2(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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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; @@ -6959,6 +6999,7 @@ x_13 = l_Lean_Compiler_compatibleTypes(x_2, x_11); 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_1); x_14 = l_Lean_indentExpr(x_4); x_15 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2___closed__2; x_16 = lean_alloc_ctor(10, 2, 0); @@ -7026,6 +7067,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_34 = !lean_is_exclusive(x_10); if (x_34 == 0) { @@ -7047,7 +7089,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2(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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; @@ -7083,6 +7125,7 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); +lean_dec(x_1); x_24 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_24, 0, x_9); lean_ctor_set(x_24, 1, x_14); @@ -7094,7 +7137,7 @@ uint8_t x_25; x_25 = !lean_is_exclusive(x_9); 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_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; x_26 = lean_ctor_get(x_9, 2); lean_dec(x_26); x_27 = lean_ctor_get(x_9, 1); @@ -7106,37 +7149,37 @@ x_30 = lean_nat_add(x_21, x_18); lean_dec(x_21); lean_ctor_set(x_9, 1, x_30); x_31 = lean_nat_dec_lt(x_6, x_3); -x_32 = lean_box(x_1); lean_inc(x_4); -x_33 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2___boxed), 9, 2); -lean_closure_set(x_33, 0, x_32); -lean_closure_set(x_33, 1, x_4); +lean_inc(x_1); +x_32 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2), 9, 2); +lean_closure_set(x_32, 0, x_1); +lean_closure_set(x_32, 1, x_4); if (x_31 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__4; -x_35 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_34); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__4; +x_34 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_33); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_36 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_35, x_29, x_33, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_36) == 0) +x_35 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_34, x_29, x_32, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_nat_add(x_6, x_8); +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_nat_add(x_6, x_8); lean_dec(x_6); x_5 = x_19; -x_6 = x_38; -x_14 = x_37; +x_6 = x_37; +x_14 = x_36; goto _start; } else { -uint8_t x_40; +uint8_t x_39; lean_dec(x_9); lean_dec(x_19); lean_dec(x_13); @@ -7145,51 +7188,52 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -x_40 = !lean_is_exclusive(x_36); -if (x_40 == 0) +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_35); +if (x_39 == 0) { -return x_36; +return x_35; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_36, 0); -x_42 = lean_ctor_get(x_36, 1); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_35, 0); +x_41 = lean_ctor_get(x_35, 1); lean_inc(x_41); -lean_dec(x_36); -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_inc(x_40); +lean_dec(x_35); +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_44; lean_object* x_45; -x_44 = lean_array_fget(x_2, x_6); +lean_object* x_43; lean_object* x_44; +x_43 = lean_array_fget(x_2, x_6); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_45 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_44, x_29, x_33, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_45) == 0) +x_44 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_43, x_29, x_32, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_44) == 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_nat_add(x_6, x_8); +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_nat_add(x_6, x_8); lean_dec(x_6); x_5 = x_19; -x_6 = x_47; -x_14 = x_46; +x_6 = x_46; +x_14 = x_45; goto _start; } else { -uint8_t x_49; +uint8_t x_48; lean_dec(x_9); lean_dec(x_19); lean_dec(x_13); @@ -7198,72 +7242,73 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -x_49 = !lean_is_exclusive(x_45); -if (x_49 == 0) +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_44); +if (x_48 == 0) { -return x_45; +return x_44; } 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_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_44, 0); +x_50 = lean_ctor_get(x_44, 1); 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_inc(x_49); +lean_dec(x_44); +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_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_dec(x_9); -x_53 = lean_array_fget(x_20, x_21); -x_54 = lean_nat_add(x_21, x_18); +x_52 = lean_array_fget(x_20, x_21); +x_53 = lean_nat_add(x_21, x_18); lean_dec(x_21); -x_55 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_55, 0, x_20); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_22); -x_56 = lean_nat_dec_lt(x_6, x_3); -x_57 = lean_box(x_1); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_20); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_22); +x_55 = lean_nat_dec_lt(x_6, x_3); lean_inc(x_4); -x_58 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2___boxed), 9, 2); -lean_closure_set(x_58, 0, x_57); -lean_closure_set(x_58, 1, x_4); -if (x_56 == 0) +lean_inc(x_1); +x_56 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2), 9, 2); +lean_closure_set(x_56, 0, x_1); +lean_closure_set(x_56, 1, x_4); +if (x_55 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__4; -x_60 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_59); +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkApp___spec__2___closed__4; +x_58 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_57); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_61 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_60, x_53, x_58, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_61) == 0) +x_59 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_58, x_52, x_56, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = lean_nat_add(x_6, x_8); +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = lean_nat_add(x_6, x_8); lean_dec(x_6); x_5 = x_19; -x_6 = x_63; -x_9 = x_55; -x_14 = x_62; +x_6 = x_61; +x_9 = x_54; +x_14 = x_60; goto _start; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -lean_dec(x_55); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_54); lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); @@ -7271,55 +7316,56 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -x_65 = lean_ctor_get(x_61, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_61, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_67 = x_61; +lean_dec(x_1); +x_63 = lean_ctor_get(x_59, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_65 = x_59; } else { - lean_dec_ref(x_61); - x_67 = lean_box(0); + lean_dec_ref(x_59); + x_65 = lean_box(0); } -if (lean_is_scalar(x_67)) { - x_68 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(1, 2, 0); } else { - x_68 = x_67; + x_66 = x_65; } -lean_ctor_set(x_68, 0, x_65); -lean_ctor_set(x_68, 1, x_66); -return x_68; +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } else { -lean_object* x_69; lean_object* x_70; -x_69 = lean_array_fget(x_2, x_6); +lean_object* x_67; lean_object* x_68; +x_67 = lean_array_fget(x_2, x_6); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_70 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_69, x_53, x_58, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_70) == 0) -{ -lean_object* x_71; lean_object* x_72; -x_71 = lean_ctor_get(x_70, 1); -lean_inc(x_71); -lean_dec(x_70); -x_72 = lean_nat_add(x_6, x_8); +x_68 = l_Lean_Compiler_lambdaBoundedTelescope___rarg(x_67, x_52, x_56, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = lean_nat_add(x_6, x_8); lean_dec(x_6); x_5 = x_19; -x_6 = x_72; -x_9 = x_55; -x_14 = x_71; +x_6 = x_70; +x_9 = x_54; +x_14 = x_69; goto _start; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_55); +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_54); lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); @@ -7327,26 +7373,27 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -x_74 = lean_ctor_get(x_70, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_70, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_70)) { - lean_ctor_release(x_70, 0); - lean_ctor_release(x_70, 1); - x_76 = x_70; +lean_dec(x_1); +x_72 = lean_ctor_get(x_68, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_68, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_74 = x_68; } else { - lean_dec_ref(x_70); - x_76 = lean_box(0); + lean_dec_ref(x_68); + x_74 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(1, 2, 0); } else { - x_77 = x_76; + x_75 = x_74; } -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_75, 0, x_72); +lean_ctor_set(x_75, 1, x_73); +return x_75; } } } @@ -7354,7 +7401,7 @@ return x_77; } else { -lean_object* x_78; +lean_object* x_76; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7362,15 +7409,16 @@ lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_9); -lean_ctor_set(x_78, 1, x_14); -return x_78; +lean_dec(x_1); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_9); +lean_ctor_set(x_76, 1, x_14); +return x_76; } } else { -lean_object* x_79; +lean_object* x_77; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7378,14 +7426,15 @@ lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_9); -lean_ctor_set(x_79, 1, x_14); -return x_79; +lean_dec(x_1); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_9); +lean_ctor_set(x_77, 1, x_14); +return x_77; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkCases(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_EXPORT lean_object* l_Lean_Compiler_check_checkCases(lean_object* x_1, lean_object* x_2, 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; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; @@ -7515,6 +7564,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_39 = !lean_is_exclusive(x_17); if (x_39 == 0) { @@ -7537,16 +7587,6 @@ return x_42; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkLambda___boxed(lean_object* x_1, lean_object* x_2, 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_1); -lean_dec(x_1); -x_10 = l_Lean_Compiler_check_checkLambda(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_check_checkBlock___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: { @@ -7602,67 +7642,23 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_3); -lean_dec(x_3); -x_11 = l_Lean_Compiler_check_checkBlock___lambda__1(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = l_Lean_Compiler_check_checkBlock___lambda__2(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_12; +x_12 = l_Lean_Compiler_check_checkBlock___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_6); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_5); -lean_dec(x_5); -x_14 = l_Lean_Compiler_check_checkBlock___lambda__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___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: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_5); -lean_dec(x_5); -x_14 = l_Lean_Compiler_check_checkBlock___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); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkBlock___boxed(lean_object* x_1, lean_object* x_2, 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_1); -lean_dec(x_1); -x_10 = l_Lean_Compiler_check_checkBlock(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkValue___boxed(lean_object* x_1, lean_object* x_2, 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_1); -lean_dec(x_1); -x_10 = lean_unbox(x_3); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); lean_dec(x_3); -x_11 = l_Lean_Compiler_check_checkValue(x_9, x_2, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; +x_10 = l_Lean_Compiler_check_checkValue(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_check_checkApp___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) { @@ -7739,36 +7735,30 @@ return x_15; LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___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: { -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_1); -lean_dec(x_1); -x_11 = l_Lean_Compiler_check_checkApp___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_10; +x_10 = l_Lean_Compiler_check_checkApp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_3); -return x_11; +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___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: { -uint8_t x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = lean_unbox(x_5); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); lean_dec(x_5); -x_14 = l_Lean_Compiler_check_checkApp___lambda__2(x_1, x_2, x_3, x_12, x_13, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; +x_13 = l_Lean_Compiler_check_checkApp___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkApp___boxed(lean_object* x_1, lean_object* x_2, 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_1); -lean_dec(x_1); -x_10 = lean_unbox(x_3); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); lean_dec(x_3); -x_11 = l_Lean_Compiler_check_checkApp(x_9, x_2, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; +x_10 = l_Lean_Compiler_check_checkApp(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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) { @@ -7798,52 +7788,28 @@ lean_dec(x_1); return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___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: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_1); -lean_dec(x_1); -x_10 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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_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_1); -lean_dec(x_1); -x_11 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2___lambda__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___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, 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_1); -lean_dec(x_1); -x_16 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__2(x_15, 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_object* x_15; +x_15 = l_Std_Range_forIn_loop___at_Lean_Compiler_check_checkCases___spec__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_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); -return x_16; +return x_15; } } LEAN_EXPORT lean_object* l_Lean_Compiler_check_checkCases___boxed(lean_object* x_1, lean_object* x_2, 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_1); -lean_dec(x_1); -x_10 = l_Lean_Compiler_check_checkCases(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_9; +x_9 = l_Lean_Compiler_check_checkCases(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); -return x_10; +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_check(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_EXPORT lean_object* l_Lean_Compiler_check(lean_object* x_1, 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; @@ -7856,12 +7822,10 @@ return x_10; LEAN_EXPORT lean_object* l_Lean_Compiler_check___boxed(lean_object* x_1, 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 = lean_unbox(x_2); -lean_dec(x_2); -x_9 = l_Lean_Compiler_check(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_object* x_8; +x_8 = l_Lean_Compiler_check(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_3); -return x_9; +return x_8; } } lean_object* initialize_Init(uint8_t builtin, lean_object*); @@ -7881,6 +7845,7 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_Util(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Compiler_Check_Config_checkJpScope___default = _init_l_Lean_Compiler_Check_Config_checkJpScope___default(); l_Lean_Compiler_Check_Config_terminalCasesOnly___default = _init_l_Lean_Compiler_Check_Config_terminalCasesOnly___default(); l_Lean_Compiler_Check_Context_jps___default = _init_l_Lean_Compiler_Check_Context_jps___default(); lean_mark_persistent(l_Lean_Compiler_Check_Context_jps___default); diff --git a/stage0/stdlib/Lean/Compiler/CompilerM.c b/stage0/stdlib/Lean/Compiler/CompilerM.c index f57bb848b0a3..68a4912cc632 100644 --- a/stage0/stdlib/Lean/Compiler/CompilerM.c +++ b/stage0/stdlib/Lean/Compiler/CompilerM.c @@ -13,23 +13,21 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3___boxed(lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_mkJump___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_mkJump___spec__5___closed__1; -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__7(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_Lean_Compiler_instVisitLetCompilerM; LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_mkJump___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_InferType_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__1(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_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___rarg(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_insert___at_Lean_Compiler_attachJp_visitLet___spec__5(lean_object*, lean_object*, lean_object*); @@ -38,12 +36,13 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_instMonadInferTypeCompilerM___boxed(lea lean_object* l_Std_HashMap_insert___at_Lean_instantiateExprMVars___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__4; -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_CompilerM_instInhabitedState; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLetUsingScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Compiler_CompilerM_instInhabitedState___closed__4; LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Compiler_attachJp_visitLet___spec__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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*); +static lean_object* l_Lean_Compiler_CompilerM_instInhabitedState___closed__1; static lean_object* l_Lean_Compiler_mkJump___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__2; static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__6; @@ -51,36 +50,35 @@ lean_object* l_Lean_Compiler_isLcCast_x3f(lean_object*); lean_object* lean_local_ctx_mk_let_decl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Compiler_compatibleTypes(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkJpDecl(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_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNames___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_EXPORT lean_object* l_Lean_Compiler_CompilerM_State_nextIdx___default; uint8_t l_Lean_Compiler_isLcUnreachable(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___boxed(lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6(lean_object*, lean_object*); -static lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__4(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_EXPORT lean_object* l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_mkLocalDecl___spec__2___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLetDecl___boxed(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_Compiler_ensureUniqueLetVarNamesCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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_Lean_mkFreshId___at_Lean_Compiler_mkLocalDecl___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_CompilerM_instInhabitedState___closed__3; lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_Compiler_instMonadLCtxCompilerM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_max(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScopeImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_CompilerM_instInhabitedState___closed__5; +static lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2; static lean_object* l_Lean_Compiler_attachJp_visitLet___closed__1; static lean_object* l_panic___at_Lean_Compiler_attachJp_visitCases___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___lambda__2(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_mkFreshId___at_Lean_Compiler_mkLocalDecl___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachOptJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,19 +90,23 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_mkLambda___boxed(lean_object*, lean_obj lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getMaxLetVarIdx___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_attachJp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__3; lean_object* l_StateRefT_x27_lift___rarg___boxed(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_Compiler_ensureUniqueLetVarNamesCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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_EXPORT lean_object* l_Lean_Compiler_instMonadLCtxCompilerM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_RBNode_setBlack___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___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_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDeclName___boxed(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_Compiler_mkFreshLetVarIdx___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1; static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3___closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_mkJump___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3___closed__3; lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -115,43 +117,47 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_instMonadLCtxCompilerM; static lean_object* l_Lean_Compiler_instAddMessageContextCompilerM___closed__3; lean_object* l_Lean_Expr_headBeta(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_mkJump___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__7___boxed(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_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___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_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLetUsingScope___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getMaxLetVarIdx(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Std_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_mkJump___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabited___rarg(lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); static lean_object* l_Lean_Compiler_attachJp_visitCases___closed__1; +static size_t l_Lean_Compiler_CompilerM_instInhabitedState___closed__2; +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); lean_object* l_Lean_Compiler_CasesInfo_updateResultingType(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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_Compiler_attachJp_visitCases___closed__2; lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___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*); static lean_object* l_Lean_Compiler_attachJp_visitLet___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8(lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at_Lean_Compiler_getMaxLetVarIdx___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_mkFreshLetVarIdx(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_instMonadLCtxCompilerM___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_instAddMessageContextCompilerM___closed__4; -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_mkJump___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instVisitLetReaderT(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instAddMessageContextCompilerM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope(lean_object*, lean_object*); @@ -162,23 +168,20 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___ LEAN_EXPORT lean_object* l_Lean_Compiler_CompilerM_State_lctx___default; static lean_object* l_Lean_Compiler_mkLetUsingScope___closed__3; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___closed__1; -static lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instAddMessageContextCompilerM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLcCast___at_Lean_Compiler_mkJump___spec__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_mkJump___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_instAddMessageContextCompilerM___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcCast___at_Lean_Compiler_mkJump___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Compiler_mkLocalDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_instMonadLCtxCompilerM___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_mk_local_decl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -190,13 +193,12 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visi lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_mkJump(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_withNewScope___rarg___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___spec__2(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_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDeclName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkJpDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Compiler_isJpBinderName(lean_object*); static lean_object* l_Lean_Compiler_instAddMessageContextCompilerM___closed__2; @@ -208,10 +210,10 @@ static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__1; lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadCoreM; -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkJpDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_mkLocalDecl___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__1; static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkJump___closed__2; @@ -220,7 +222,6 @@ lean_object* l_Lean_LocalContext_mkLambda(lean_object*, lean_object*, lean_objec uint8_t l_Lean_LocalDecl_isJp(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__2(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_isAnyType(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(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_set(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -229,35 +230,39 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_mkJpDeclIfNotSimple___boxed(lean_object LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLcCast___at_Lean_Compiler_mkJump___spec__2___closed__1; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instVisitLetStateRefT_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getMaxLetVarIdx___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__2(lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___spec__4(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_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLetUsingScope___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_375_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_mkJpDeclIfNotSimple(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_instVisitLetCompilerM___closed__1; LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at_Lean_Compiler_getMaxLetVarIdx___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CompilerM_State_lctx___default___closed__7; lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1; lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_attachJp_visitCases___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLocalDecl(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CompilerM_State_letFVars___default; lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNames(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___spec__7(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_EXPORT lean_object* l_Lean_Compiler_visitMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -269,17 +274,17 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_instVisitLetStateRefT_x27___rarg(lean_o lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); lean_object* l_Std_HashMap_insert___at_Lean_ForEachExpr_visit___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instMonadLCtxCompilerM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_attachJp___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore_go___boxed(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_ensureUniqueLetVarNamesCore___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_instMonadInferTypeCompilerM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcCast___at_Lean_Compiler_mkJump___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_mkJump___spec__1___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_attachJp_visitLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Compiler_CompilerM_State_lctx___default___closed__1() { _start: { @@ -391,6 +396,76 @@ x_1 = lean_unsigned_to_nat(1u); return x_1; } } +static lean_object* _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static size_t _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__2() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_usize_of_nat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; size_t x_4; lean_object* x_5; +x_1 = l_Lean_Compiler_CompilerM_instInhabitedState___closed__1; +x_2 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Lean_Compiler_CompilerM_instInhabitedState___closed__2; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*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(x_5, 3, x_3); +lean_ctor_set_usize(x_5, 4, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_CompilerM_State_lctx___default___closed__3; +x_2 = l_Lean_Compiler_CompilerM_instInhabitedState___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_Compiler_CompilerM_instInhabitedState___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Compiler_CompilerM_instInhabitedState___closed__4; +x_2 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_3 = lean_unsigned_to_nat(0u); +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_Lean_Compiler_CompilerM_instInhabitedState() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_CompilerM_instInhabitedState___closed__5; +return x_1; +} +} static lean_object* _init_l_Lean_Compiler_instAddMessageContextCompilerM___closed__1() { _start: { @@ -1166,327 +1241,222 @@ lean_dec(x_5); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDecl(lean_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_Compiler_mkFreshLetVarIdx(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; uint8_t x_20; -x_20 = l_Lean_Expr_isFVar(x_1); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; 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_48 = lean_st_ref_get(x_5, x_6); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = lean_st_ref_get(x_3, x_49); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = lean_st_ref_get(x_5, x_52); -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -lean_dec(x_53); -x_55 = lean_st_ref_get(x_3, x_54); -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); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_59 = l_Lean_Compiler_InferType_inferType(x_1, x_58, x_4, x_5, x_57); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t 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; uint8_t x_73; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_51, 2); -lean_inc(x_62); -lean_dec(x_51); -x_63 = l_Lean_Name_num___override(x_2, x_62); -x_64 = 0; -x_65 = l_Lean_Compiler_mkLetDecl(x_63, x_60, x_1, x_64, x_3, x_4, x_5, x_61); -lean_dec(x_4); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_st_ref_get(x_5, x_67); +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_5 = lean_st_ref_get(x_3, x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); lean_dec(x_5); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = lean_st_ref_take(x_3, 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); -lean_dec(x_70); -x_73 = !lean_is_exclusive(x_71); -if (x_73 == 0) +x_7 = lean_st_ref_take(x_1, x_6); +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_is_exclusive(x_8); +if (x_10 == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_74 = lean_ctor_get(x_71, 2); -x_75 = lean_unsigned_to_nat(1u); -x_76 = lean_nat_add(x_74, x_75); -lean_dec(x_74); -lean_ctor_set(x_71, 2, x_76); -x_77 = lean_st_ref_set(x_3, x_71, x_72); -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 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_8, 2); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_11, x_12); +lean_ctor_set(x_8, 2, x_13); +x_14 = lean_st_ref_set(x_1, x_8, x_9); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_77, 0); -lean_dec(x_79); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_66); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_77, 0, x_81); -x_7 = x_77; -goto block_19; +lean_object* x_16; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_11); +return x_14; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_82 = lean_ctor_get(x_77, 1); -lean_inc(x_82); -lean_dec(x_77); -x_83 = lean_box(0); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_66); -lean_ctor_set(x_84, 1, x_83); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -x_7 = x_85; -goto block_19; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } 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; -x_86 = lean_ctor_get(x_71, 0); -x_87 = lean_ctor_get(x_71, 1); -x_88 = lean_ctor_get(x_71, 2); -lean_inc(x_88); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_71); -x_89 = lean_unsigned_to_nat(1u); -x_90 = lean_nat_add(x_88, x_89); -lean_dec(x_88); -x_91 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_91, 0, x_86); -lean_ctor_set(x_91, 1, x_87); -lean_ctor_set(x_91, 2, x_90); -x_92 = lean_st_ref_set(x_3, x_91, x_72); -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_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 = lean_ctor_get(x_8, 0); +x_20 = lean_ctor_get(x_8, 1); +x_21 = lean_ctor_get(x_8, 2); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_8); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_21, x_22); +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_20); +lean_ctor_set(x_24, 2, x_23); +x_25 = lean_st_ref_set(x_1, x_24, x_9); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_27 = x_25; } else { - lean_dec_ref(x_92); - x_94 = lean_box(0); + lean_dec_ref(x_25); + x_27 = lean_box(0); } -x_95 = lean_box(0); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_66); -lean_ctor_set(x_96, 1, x_95); -if (lean_is_scalar(x_94)) { - x_97 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_27)) { + x_28 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_94; + x_28 = x_27; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_93); -x_7 = x_97; -goto block_19; +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_mkFreshLetVarIdx___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_98; lean_object* x_99; -lean_dec(x_51); -lean_dec(x_4); +lean_object* x_5; +x_5 = l_Lean_Compiler_mkFreshLetVarIdx(x_1, x_2, x_3, x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_98 = lean_ctor_get(x_59, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_59, 1); -lean_inc(x_99); -lean_dec(x_59); -x_21 = x_98; -x_22 = x_99; -goto block_47; +return x_5; } -block_47: -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_23 = lean_st_ref_get(x_5, x_22); -lean_dec(x_5); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_st_ref_take(x_3, x_24); -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_is_exclusive(x_26); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_29 = lean_ctor_get(x_26, 2); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_29, x_30); -lean_dec(x_29); -lean_ctor_set(x_26, 2, x_31); -x_32 = lean_st_ref_set(x_3, x_26, x_27); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -lean_ctor_set_tag(x_32, 1); -lean_ctor_set(x_32, 0, x_21); -x_7 = x_32; -goto block_19; } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDeclName(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_21); -lean_ctor_set(x_36, 1, x_35); -x_7 = x_36; -goto block_19; -} +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_Compiler_mkFreshLetVarIdx(x_2, x_3, x_4, x_5); +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, 0); +x_9 = l_Lean_Name_num___override(x_1, x_8); +lean_ctor_set(x_6, 0, x_9); +return x_6; } 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; lean_object* x_45; lean_object* x_46; -x_37 = lean_ctor_get(x_26, 0); -x_38 = lean_ctor_get(x_26, 1); -x_39 = lean_ctor_get(x_26, 2); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_26); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_add(x_39, x_40); -lean_dec(x_39); -x_42 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_42, 0, x_37); -lean_ctor_set(x_42, 1, x_38); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_st_ref_set(x_3, x_42, x_27); -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -if (lean_is_exclusive(x_43)) { - lean_ctor_release(x_43, 0); - lean_ctor_release(x_43, 1); - x_45 = x_43; -} else { - lean_dec_ref(x_43); - x_45 = lean_box(0); -} -if (lean_is_scalar(x_45)) { - x_46 = lean_alloc_ctor(1, 2, 0); -} else { - x_46 = x_45; - lean_ctor_set_tag(x_46, 1); -} -lean_ctor_set(x_46, 0, x_21); -lean_ctor_set(x_46, 1, x_44); -x_7 = x_46; -goto block_19; +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_Lean_Name_num___override(x_1, x_10); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; } } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDeclName___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_100; -lean_dec(x_5); +lean_object* x_6; +x_6 = l_Lean_Compiler_mkAuxLetDeclName(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_1); -lean_ctor_set(x_100, 1, x_6); -return x_100; +return x_6; } -block_19: -{ -if (lean_obj_tag(x_7) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Compiler_mkAuxLetDecl(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_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +uint8_t x_7; +x_7 = l_Lean_Expr_isFVar(x_1); +if (x_7 == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_9, 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; +x_8 = l_Lean_Compiler_mkAuxLetDeclName(x_2, x_3, x_4, x_5, 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_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); +lean_dec(x_8); +x_11 = lean_st_ref_get(x_5, x_10); +x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); lean_dec(x_11); -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 -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) +x_13 = lean_st_ref_get(x_3, 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_14, 0); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_17 = l_Lean_Compiler_InferType_inferType(x_1, x_16, x_4, x_5, x_15); +if (lean_obj_tag(x_17) == 0) { -return x_7; +lean_object* x_18; lean_object* x_19; uint8_t 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 = 0; +x_21 = l_Lean_Compiler_mkLetDecl(x_9, x_18, x_1, x_20, x_3, x_4, x_5, x_19); +lean_dec(x_5); +lean_dec(x_4); +return x_21; } 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; +uint8_t x_22; +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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_5); +lean_dec(x_4); +lean_dec(x_2); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_6); +return x_26; } } } @@ -2907,7 +2877,7 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___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) { +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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: { lean_object* x_11; @@ -2968,7 +2938,7 @@ lean_dec(x_11); x_20 = lean_ctor_get(x_12, 0); lean_inc(x_20); lean_dec(x_12); -x_21 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_20, x_6, x_7, x_8, x_9, x_19); +x_21 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_20, x_6, x_7, x_8, x_9, x_19); return x_21; } } @@ -3004,7 +2974,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___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, 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_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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, 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; @@ -3039,7 +3009,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_18 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_15, x_8, x_9, x_10, x_11, x_12); +x_18 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_15, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; @@ -3090,7 +3060,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNames___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, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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, lean_object* x_11, lean_object* x_12) { _start: { if (lean_obj_tag(x_5) == 5) @@ -3122,7 +3092,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_19 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +x_19 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; @@ -3143,7 +3113,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___spec__4(x_1, x_2, x_3, x_4, x_23, x_24, x_6, x_8, x_9, x_10, x_11, x_21); +x_25 = l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__4(x_1, x_2, x_3, x_4, x_23, x_24, x_6, x_8, x_9, x_10, x_11, x_21); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -3153,7 +3123,7 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_mkAppN(x_20, x_26); -x_29 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_28, x_8, x_9, x_10, x_11, x_27); +x_29 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_28, x_8, x_9, x_10, x_11, x_27); return x_29; } else @@ -3222,7 +3192,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___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_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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) { _start: { lean_object* x_8; @@ -3271,15 +3241,15 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6___rarg), 7, 0); +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6___rarg), 7, 0); return x_3; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3289,21 +3259,21 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___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_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg(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_Compiler_ensureUniqueLetVarNamesCore___spec__8___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 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2; +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -3313,15 +3283,15 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___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_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -3513,7 +3483,7 @@ lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_55 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg(x_13, x_5, x_6, x_7); +x_55 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg(x_13, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); x_56 = !lean_is_exclusive(x_55); @@ -3537,7 +3507,7 @@ return x_59; } } } -static lean_object* _init_l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3546,7 +3516,7 @@ x_2 = l_Lean_Expr_sort___override(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___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_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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) { _start: { if (lean_obj_tag(x_5) == 0) @@ -3580,13 +3550,13 @@ case 5: 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_unsigned_to_nat(0u); x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_13, x_14); -x_16 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1; +x_16 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1; lean_inc(x_15); x_17 = lean_mk_array(x_15, x_16); x_18 = lean_unsigned_to_nat(1u); x_19 = lean_nat_sub(x_15, x_18); lean_dec(x_15); -x_20 = l_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNames___spec__5(x_1, x_2, x_3, x_4, x_13, x_17, x_19, x_6, x_7, x_8, x_9, x_10); +x_20 = l_Lean_Expr_withAppAux___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__5(x_1, x_2, x_3, x_4, x_13, x_17, x_19, x_6, x_7, x_8, x_9, x_10); return x_20; } case 6: @@ -3609,7 +3579,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_25 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_22, x_6, x_7, x_8, x_9, x_10); +x_25 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_22, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -3627,7 +3597,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_28 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_23, x_6, x_7, x_8, x_9, x_27); +x_28 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_23, x_6, x_7, x_8, x_9, x_27); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; uint8_t x_34; @@ -3650,7 +3620,7 @@ lean_object* x_35; lean_object* x_36; lean_dec(x_31); lean_dec(x_23); x_35 = l_Lean_Expr_lam___override(x_21, x_26, x_29, x_24); -x_36 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_35, x_6, x_7, x_8, x_9, x_30); +x_36 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_35, x_6, x_7, x_8, x_9, x_30); return x_36; } else @@ -3665,7 +3635,7 @@ if (x_39 == 0) lean_object* x_40; lean_object* x_41; lean_dec(x_31); x_40 = l_Lean_Expr_lam___override(x_21, x_26, x_29, x_24); -x_41 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_40, x_6, x_7, x_8, x_9, x_30); +x_41 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_40, x_6, x_7, x_8, x_9, x_30); return x_41; } else @@ -3677,7 +3647,7 @@ if (x_42 == 0) lean_object* x_43; lean_object* x_44; lean_dec(x_31); x_43 = l_Lean_Expr_lam___override(x_21, x_26, x_29, x_24); -x_44 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_43, x_6, x_7, x_8, x_9, x_30); +x_44 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_43, x_6, x_7, x_8, x_9, x_30); return x_44; } else @@ -3686,7 +3656,7 @@ lean_object* x_45; lean_dec(x_29); lean_dec(x_26); lean_dec(x_21); -x_45 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_31, x_6, x_7, x_8, x_9, x_30); +x_45 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_31, x_6, x_7, x_8, x_9, x_30); return x_45; } } @@ -3781,7 +3751,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_58 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_55, x_6, x_7, x_8, x_9, x_10); +x_58 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_55, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_58) == 0) { lean_object* x_59; lean_object* x_60; lean_object* x_61; @@ -3799,7 +3769,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_61 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_56, x_6, x_7, x_8, x_9, x_60); +x_61 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_56, x_6, x_7, x_8, x_9, x_60); if (lean_obj_tag(x_61) == 0) { lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; uint8_t x_67; @@ -3822,7 +3792,7 @@ lean_object* x_68; lean_object* x_69; lean_dec(x_64); lean_dec(x_56); x_68 = l_Lean_Expr_forallE___override(x_54, x_59, x_62, x_57); -x_69 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_68, x_6, x_7, x_8, x_9, x_63); +x_69 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_68, x_6, x_7, x_8, x_9, x_63); return x_69; } else @@ -3837,7 +3807,7 @@ if (x_72 == 0) lean_object* x_73; lean_object* x_74; lean_dec(x_64); x_73 = l_Lean_Expr_forallE___override(x_54, x_59, x_62, x_57); -x_74 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_73, x_6, x_7, x_8, x_9, x_63); +x_74 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_73, x_6, x_7, x_8, x_9, x_63); return x_74; } else @@ -3849,7 +3819,7 @@ if (x_75 == 0) lean_object* x_76; lean_object* x_77; lean_dec(x_64); x_76 = l_Lean_Expr_forallE___override(x_54, x_59, x_62, x_57); -x_77 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_76, x_6, x_7, x_8, x_9, x_63); +x_77 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_76, x_6, x_7, x_8, x_9, x_63); return x_77; } else @@ -3858,7 +3828,7 @@ lean_object* x_78; lean_dec(x_62); lean_dec(x_59); lean_dec(x_54); -x_78 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_64, x_6, x_7, x_8, x_9, x_63); +x_78 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_64, x_6, x_7, x_8, x_9, x_63); return x_78; } } @@ -3954,7 +3924,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_92 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_88, x_6, x_7, x_8, x_9, x_10); +x_92 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_88, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_92) == 0) { lean_object* x_93; lean_object* x_94; lean_object* x_95; @@ -3972,7 +3942,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_95 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_89, x_6, x_7, x_8, x_9, x_94); +x_95 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_89, x_6, x_7, x_8, x_9, x_94); if (lean_obj_tag(x_95) == 0) { lean_object* x_96; lean_object* x_97; lean_object* x_98; @@ -3990,7 +3960,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_98 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_90, x_6, x_7, x_8, x_9, x_97); +x_98 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_90, x_6, x_7, x_8, x_9, x_97); if (lean_obj_tag(x_98) == 0) { lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; uint8_t x_103; @@ -4010,7 +3980,7 @@ lean_dec(x_90); lean_dec(x_89); lean_dec(x_13); x_104 = l_Lean_Expr_letE___override(x_87, x_93, x_96, x_99, x_91); -x_105 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_104, x_6, x_7, x_8, x_9, x_100); +x_105 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_104, x_6, x_7, x_8, x_9, x_100); return x_105; } else @@ -4026,7 +3996,7 @@ lean_object* x_109; lean_object* x_110; lean_dec(x_90); lean_dec(x_13); x_109 = l_Lean_Expr_letE___override(x_87, x_93, x_96, x_99, x_91); -x_110 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_109, x_6, x_7, x_8, x_9, x_100); +x_110 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_109, x_6, x_7, x_8, x_9, x_100); return x_110; } else @@ -4041,7 +4011,7 @@ if (x_113 == 0) lean_object* x_114; lean_object* x_115; lean_dec(x_13); x_114 = l_Lean_Expr_letE___override(x_87, x_93, x_96, x_99, x_91); -x_115 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_114, x_6, x_7, x_8, x_9, x_100); +x_115 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_114, x_6, x_7, x_8, x_9, x_100); return x_115; } else @@ -4051,7 +4021,7 @@ lean_dec(x_99); lean_dec(x_96); lean_dec(x_93); lean_dec(x_87); -x_116 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_100); +x_116 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_100); return x_116; } } @@ -4184,7 +4154,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_131 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_130, x_6, x_7, x_8, x_9, x_10); +x_131 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_130, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_131) == 0) { lean_object* x_132; lean_object* x_133; size_t x_134; size_t x_135; uint8_t x_136; @@ -4202,7 +4172,7 @@ if (x_136 == 0) lean_object* x_137; lean_object* x_138; lean_dec(x_13); x_137 = l_Lean_Expr_mdata___override(x_129, x_132); -x_138 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_137, x_6, x_7, x_8, x_9, x_133); +x_138 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_137, x_6, x_7, x_8, x_9, x_133); return x_138; } else @@ -4210,7 +4180,7 @@ else lean_object* x_139; lean_dec(x_132); lean_dec(x_129); -x_139 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_133); +x_139 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_133); return x_139; } } @@ -4266,7 +4236,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_147 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_1, x_2, x_3, x_4, x_146, x_6, x_7, x_8, x_9, x_10); +x_147 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_1, x_2, x_3, x_4, x_146, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_147) == 0) { lean_object* x_148; lean_object* x_149; size_t x_150; size_t x_151; uint8_t x_152; @@ -4284,7 +4254,7 @@ if (x_152 == 0) lean_object* x_153; lean_object* x_154; lean_dec(x_13); x_153 = l_Lean_Expr_proj___override(x_144, x_145, x_148); -x_154 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_153, x_6, x_7, x_8, x_9, x_149); +x_154 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_153, x_6, x_7, x_8, x_9, x_149); return x_154; } else @@ -4293,7 +4263,7 @@ lean_object* x_155; lean_dec(x_148); lean_dec(x_145); lean_dec(x_144); -x_155 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_149); +x_155 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_149); return x_155; } } @@ -4335,14 +4305,14 @@ return x_159; default: { lean_object* x_160; -x_160 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNames___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10); +x_160 = l_Lean_Core_transform_visit_visitPost___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__3(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10); return x_160; } } } } } -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___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; @@ -4354,7 +4324,7 @@ lean_ctor_set(x_6, 1, x_4); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___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) { +LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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: { lean_object* x_11; lean_object* x_12; @@ -4390,19 +4360,19 @@ x_18 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_18, 0, x_17); lean_inc(x_4); lean_inc(x_3); -x_19 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1), 10, 4); +x_19 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1), 10, 4); lean_closure_set(x_19, 0, x_1); lean_closure_set(x_19, 1, x_2); lean_closure_set(x_19, 2, x_3); lean_closure_set(x_19, 3, x_4); -x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6___rarg), 7, 2); +x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6___rarg), 7, 2); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_19); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_21 = l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___spec__7(x_3, x_20, x_6, x_7, x_8, x_9, x_15); +x_21 = l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__7(x_3, x_20, x_6, x_7, x_8, x_9, x_15); lean_dec(x_3); if (lean_obj_tag(x_21) == 0) { @@ -4413,7 +4383,7 @@ x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); lean_inc(x_22); -x_24 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__2), 3, 2); +x_24 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__2), 3, 2); lean_closure_set(x_24, 0, x_5); lean_closure_set(x_24, 1, x_22); x_25 = lean_alloc_closure((void*)(l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed), 3, 2); @@ -4536,19 +4506,19 @@ x_44 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_44, 0, x_43); lean_inc(x_4); lean_inc(x_3); -x_45 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1), 10, 4); +x_45 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1), 10, 4); lean_closure_set(x_45, 0, x_1); lean_closure_set(x_45, 1, x_2); lean_closure_set(x_45, 2, x_3); lean_closure_set(x_45, 3, x_4); -x_46 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNames___spec__6___rarg), 7, 2); +x_46 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__6___rarg), 7, 2); lean_closure_set(x_46, 0, x_44); lean_closure_set(x_46, 1, x_45); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_47 = l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___spec__7(x_3, x_46, x_6, x_7, x_8, x_9, x_41); +x_47 = l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__7(x_3, x_46, x_6, x_7, x_8, x_9, x_41); lean_dec(x_3); if (lean_obj_tag(x_47) == 0) { @@ -4559,7 +4529,7 @@ x_49 = lean_ctor_get(x_47, 1); lean_inc(x_49); lean_dec(x_47); lean_inc(x_48); -x_50 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__2), 3, 2); +x_50 = lean_alloc_closure((void*)(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__2), 3, 2); lean_closure_set(x_50, 0, x_5); lean_closure_set(x_50, 1, x_48); x_51 = lean_alloc_closure((void*)(l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed), 3, 2); @@ -4700,7 +4670,7 @@ return x_69; } } } -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___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_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -4729,15 +4699,15 @@ return x_13; } } } -static lean_object* _init_l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1() { +static lean_object* _init_l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___lambda__1___boxed), 6, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___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 lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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; 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; @@ -4753,10 +4723,10 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1; +x_15 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1; lean_inc(x_6); lean_inc(x_13); -x_16 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2(x_2, x_3, x_8, x_15, x_1, x_13, x_4, x_5, x_6, x_14); +x_16 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2(x_2, x_3, x_8, x_15, x_1, x_13, x_4, x_5, x_6, 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; uint8_t x_22; @@ -4819,7 +4789,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNamesCore___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_1) == 8) @@ -4931,7 +4901,7 @@ return x_42; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNamesCore___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; @@ -4943,105 +4913,33 @@ lean_ctor_set(x_7, 1, x_5); return x_7; } } -static lean_object* _init_l_Lean_Compiler_ensureUniqueLetVarNames___closed__1() { +static lean_object* _init_l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNames___lambda__1___boxed), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__1___boxed), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Compiler_ensureUniqueLetVarNames___closed__2() { +static lean_object* _init_l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNames___lambda__2___boxed), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__2___boxed), 5, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _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; -x_5 = lean_st_ref_get(x_3, x_4); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_unsigned_to_nat(1u); -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_Lean_Compiler_ensureUniqueLetVarNames___closed__1; -x_12 = l_Lean_Compiler_ensureUniqueLetVarNames___closed__2; -lean_inc(x_3); -lean_inc(x_9); -x_13 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1(x_1, x_11, x_12, x_9, x_2, x_3, x_10); -if (lean_obj_tag(x_13) == 0) -{ -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_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_st_ref_get(x_3, x_15); -lean_dec(x_3); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_st_ref_get(x_9, x_17); -lean_dec(x_9); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_14); -return x_18; -} -else -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -else -{ -uint8_t x_23; -lean_dec(x_9); -lean_dec(x_3); -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; -} -else -{ -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; -} -} +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1; +x_7 = l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2; +x_8 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(x_1, x_6, x_7, x_2, x_3, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___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, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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, 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; @@ -5049,72 +4947,231 @@ x_13 = lean_unbox_usize(x_5); lean_dec(x_5); x_14 = lean_unbox_usize(x_6); lean_dec(x_6); -x_15 = l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNames___spec__4(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__4(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___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_Compiler_ensureUniqueLetVarNamesCore___spec__8___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_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg(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_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8(x_1); +x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___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_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNames___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Core_withIncRecDepth___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___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_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___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) { _start: { lean_object* x_7; -x_7 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___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_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNamesCore___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_Compiler_ensureUniqueLetVarNames___lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Compiler_ensureUniqueLetVarNamesCore___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_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNamesCore___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_Compiler_ensureUniqueLetVarNames___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__2(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_Compiler_visitLambda_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_Lean_Compiler_ensureUniqueLetVarNames(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; +x_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_st_ref_get(x_2, 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, 2); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_st_ref_get(x_4, x_10); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_mk_ref(x_11, x_13); +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_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1; +x_18 = l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2; +lean_inc(x_4); +lean_inc(x_15); +x_19 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(x_1, x_17, x_18, x_15, x_3, x_4, x_16); +if (lean_obj_tag(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; uint8_t x_32; +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_st_ref_get(x_4, x_21); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_st_ref_get(x_15, x_23); +lean_dec(x_15); +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_st_ref_get(x_4, x_26); +lean_dec(x_4); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_st_ref_take(x_2, x_28); +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 = !lean_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_30, 2); +lean_dec(x_33); +lean_ctor_set(x_30, 2, x_25); +x_34 = lean_st_ref_set(x_2, x_30, x_31); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_20); +return x_34; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_20); +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; +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(0, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_25); +x_42 = lean_st_ref_set(x_2, x_41, x_31); +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); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_20); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} +} +else +{ +uint8_t x_46; +lean_dec(x_15); +lean_dec(x_4); +x_46 = !lean_is_exclusive(x_19); +if (x_46 == 0) +{ +return x_19; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_19, 0); +x_48 = lean_ctor_get(x_19, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_19); +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_EXPORT lean_object* l_Lean_Compiler_ensureUniqueLetVarNames___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_Compiler_ensureUniqueLetVarNames(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore_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_1) == 6) @@ -5144,75 +5201,158 @@ goto _start; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_expr_instantiate_rev(x_1, x_2); -lean_dec(x_1); +lean_object* x_17; lean_object* x_18; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_1); x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_2); -lean_ctor_set(x_18, 1, 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_6); -return x_19; +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_6); +return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda_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_EXPORT lean_object* l_Lean_Compiler_visitLambdaCore_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_Compiler_visitLambda_go(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Compiler_visitLambdaCore_go(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_Compiler_visitLambda(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_Compiler_visitLambdaCore(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 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; -x_7 = l_Lean_Compiler_visitLambda_go(x_1, x_6, x_2, x_3, x_4, x_5); +x_7 = l_Lean_Compiler_visitLambdaCore_go(x_1, x_6, x_2, x_3, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda___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_Compiler_visitLambdaCore___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_Compiler_visitLambda(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Compiler_visitLambdaCore(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; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLambda(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_mk_string_from_bytes("Init.Util", 9); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__2() { -_start: +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_Compiler_visitLambdaCore(x_1, x_2, x_3, x_4, x_5); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("getElem!", 8); -return x_1; -} -} -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__3() { -_start: +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_1; -x_1 = lean_mk_string_from_bytes("index out of bounds", 19); -return x_1; -} +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); +x_12 = lean_expr_instantiate_rev(x_11, x_10); +lean_dec(x_11); +lean_ctor_set(x_8, 1, x_12); +return x_6; } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__4() { +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +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 = lean_expr_instantiate_rev(x_14, x_13); +lean_dec(x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_6, 0, x_16); +return x_6; +} +} +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; +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_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + x_21 = x_17; +} else { + lean_dec_ref(x_17); + x_21 = lean_box(0); +} +x_22 = lean_expr_instantiate_rev(x_20, x_19); +lean_dec(x_20); +if (lean_is_scalar(x_21)) { + x_23 = lean_alloc_ctor(0, 2, 0); +} else { + x_23 = x_21; +} +lean_ctor_set(x_23, 0, x_19); +lean_ctor_set(x_23, 1, 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_Compiler_visitLambda___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_Compiler_visitLambda(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; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Init.Util", 9); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("getElem!", 8); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("index out of bounds", 19); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___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; @@ -5387,7 +5527,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_visitMatch(lean_object* x_1, lean_objec lean_object* x_7; lean_object* x_8; 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; uint8_t x_30; x_7 = lean_unsigned_to_nat(0u); x_8 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_7); -x_9 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1; +x_9 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1; lean_inc(x_8); x_10 = lean_mk_array(x_8, x_9); x_11 = lean_unsigned_to_nat(1u); @@ -5916,155 +6056,545 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___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) { _start: { -lean_object* x_7; -x_7 = l_Lean_Compiler_withNewScopeImp___rarg(x_2, x_3, x_4, x_5, x_6); -return x_7; -} -} -static lean_object* _init_l_Lean_Compiler_withNewScope___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_withNewScope___rarg___lambda__1), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_Compiler_withNewScope___rarg___closed__1; -x_4 = lean_apply_3(x_1, lean_box(0), x_3, x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_withNewScope___rarg), 2, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_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_object* x_7) { -_start: -{ -if (lean_obj_tag(x_2) == 8) -{ -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; -x_8 = lean_ctor_get(x_2, 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; uint8_t x_17; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); -x_9 = lean_ctor_get(x_2, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_2, 2); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_3, x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -x_11 = lean_ctor_get(x_2, 3); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 8); -lean_dec(x_2); -x_13 = lean_expr_instantiate_rev(x_9, x_3); lean_dec(x_9); -x_14 = lean_expr_instantiate_rev(x_10, x_3); -lean_dec(x_10); -lean_inc(x_1); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_8); -x_15 = lean_apply_6(x_1, x_8, x_14, x_4, x_5, x_6, x_7); -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); +x_12 = lean_st_ref_get(x_5, x_11); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_take(x_3, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_Compiler_mkLetDecl(x_8, x_13, x_16, x_12, x_4, x_5, x_6, 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_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 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_15, 1); lean_dec(x_18); -x_21 = lean_array_push(x_3, x_19); -x_2 = x_11; -x_3 = x_21; -x_7 = x_20; -goto _start; -} -else +x_19 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +lean_ctor_set(x_15, 1, x_19); +x_20 = lean_st_ref_set(x_3, x_15, x_16); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +lean_inc(x_5); +lean_inc(x_3); +x_22 = lean_apply_4(x_2, x_3, x_4, x_5, x_21); +if (lean_obj_tag(x_22) == 0) { -uint8_t x_23; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_6); +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; +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_st_ref_get(x_5, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_3, x_26); +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_10, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = !lean_is_exclusive(x_28); +if (x_32 == 0) +{ +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_33 = lean_ctor_get(x_28, 1); +lean_dec(x_33); +x_34 = lean_ctor_get(x_28, 0); +lean_dec(x_34); +lean_ctor_set(x_28, 1, x_31); +lean_ctor_set(x_28, 0, x_30); +x_35 = lean_st_ref_get(x_5, x_29); lean_dec(x_5); -lean_dec(x_4); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_set(x_3, x_28, x_36); lean_dec(x_3); -lean_dec(x_1); -x_23 = !lean_is_exclusive(x_15); -if (x_23 == 0) +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -return x_15; +lean_object* x_39; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +lean_ctor_set(x_37, 0, x_23); +return x_37; } 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_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_23); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } else { -lean_object* x_27; lean_object* x_28; -lean_dec(x_6); +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_42 = lean_ctor_get(x_28, 2); +lean_inc(x_42); +lean_dec(x_28); +x_43 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_43, 0, x_30); +lean_ctor_set(x_43, 1, x_31); +lean_ctor_set(x_43, 2, x_42); +x_44 = lean_st_ref_get(x_5, x_29); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_27 = lean_expr_instantiate_rev(x_2, x_3); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_st_ref_set(x_3, x_43, x_45); lean_dec(x_3); -lean_dec(x_2); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_7); -return x_28; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp(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_Compiler_CompilerM_State_letFVars___default___closed__1; -x_8 = l_Lean_Compiler_visitLetImp_go(x_2, x_1, x_7, x_3, x_4, x_5, x_6); -return x_8; +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_48 = x_46; +} else { + lean_dec_ref(x_46); + x_48 = lean_box(0); } +if (lean_is_scalar(x_48)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_48; } -static lean_object* _init_l_Lean_Compiler_instVisitLetCompilerM___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLetImp), 6, 0); -return x_1; +lean_ctor_set(x_49, 0, x_23); +lean_ctor_set(x_49, 1, x_47); +return x_49; } } -static lean_object* _init_l_Lean_Compiler_instVisitLetCompilerM() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_Compiler_instVisitLetCompilerM___closed__1; -return x_1; -} -} +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; uint8_t x_59; +x_50 = lean_ctor_get(x_22, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_22, 1); +lean_inc(x_51); +lean_dec(x_22); +x_52 = lean_st_ref_get(x_5, x_51); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_st_ref_get(x_3, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_10, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_10, 1); +lean_inc(x_58); +lean_dec(x_10); +x_59 = !lean_is_exclusive(x_55); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_60 = lean_ctor_get(x_55, 1); +lean_dec(x_60); +x_61 = lean_ctor_get(x_55, 0); +lean_dec(x_61); +lean_ctor_set(x_55, 1, x_58); +lean_ctor_set(x_55, 0, x_57); +x_62 = lean_st_ref_get(x_5, x_56); +lean_dec(x_5); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = lean_st_ref_set(x_3, x_55, x_63); +lean_dec(x_3); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +lean_ctor_set_tag(x_64, 1); +lean_ctor_set(x_64, 0, x_50); +return x_64; +} +else +{ +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_64, 1); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_50); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} +} +else +{ +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_69 = lean_ctor_get(x_55, 2); +lean_inc(x_69); +lean_dec(x_55); +x_70 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_70, 0, x_57); +lean_ctor_set(x_70, 1, x_58); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_st_ref_get(x_5, x_56); +lean_dec(x_5); +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +lean_dec(x_71); +x_73 = lean_st_ref_set(x_3, x_70, x_72); +lean_dec(x_3); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_75 = x_73; +} else { + lean_dec_ref(x_73); + 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_tag(x_76, 1); +} +lean_ctor_set(x_76, 0, x_50); +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; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_77 = lean_ctor_get(x_15, 0); +x_78 = lean_ctor_get(x_15, 2); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_15); +x_79 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_80 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_80, 2, x_78); +x_81 = lean_st_ref_set(x_3, x_80, x_16); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +lean_inc(x_5); +lean_inc(x_3); +x_83 = lean_apply_4(x_2, x_3, x_4, x_5, x_82); +if (lean_obj_tag(x_83) == 0) +{ +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_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_5, x_85); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = lean_st_ref_get(x_3, x_87); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = lean_ctor_get(x_10, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_10, 1); +lean_inc(x_92); +lean_dec(x_10); +x_93 = lean_ctor_get(x_89, 2); +lean_inc(x_93); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + lean_ctor_release(x_89, 2); + x_94 = x_89; +} else { + lean_dec_ref(x_89); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(0, 3, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_91); +lean_ctor_set(x_95, 1, x_92); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_st_ref_get(x_5, x_90); +lean_dec(x_5); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = lean_st_ref_set(x_3, x_95, x_97); +lean_dec(x_3); +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_98); + x_100 = lean_box(0); +} +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_100; +} +lean_ctor_set(x_101, 0, x_84); +lean_ctor_set(x_101, 1, x_99); +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; 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; +x_102 = lean_ctor_get(x_83, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_83, 1); +lean_inc(x_103); +lean_dec(x_83); +x_104 = lean_st_ref_get(x_5, x_103); +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_st_ref_get(x_3, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_10, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_10, 1); +lean_inc(x_110); +lean_dec(x_10); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + x_112 = x_107; +} else { + lean_dec_ref(x_107); + x_112 = lean_box(0); +} +if (lean_is_scalar(x_112)) { + x_113 = lean_alloc_ctor(0, 3, 0); +} else { + x_113 = x_112; +} +lean_ctor_set(x_113, 0, x_109); +lean_ctor_set(x_113, 1, x_110); +lean_ctor_set(x_113, 2, x_111); +x_114 = lean_st_ref_get(x_5, x_108); +lean_dec(x_5); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_set(x_3, x_113, x_115); +lean_dec(x_3); +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_118 = x_116; +} else { + lean_dec_ref(x_116); + 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_tag(x_119, 1); +} +lean_ctor_set(x_119, 0, x_102); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +} +} +} +static lean_object* _init_l_Lean_Compiler_withNewScope___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_withNewScope___rarg___lambda__1), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_Compiler_withNewScope___rarg___closed__1; +x_4 = lean_apply_3(x_1, lean_box(0), x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_withNewScope___rarg), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_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_object* x_7) { +_start: +{ +if (lean_obj_tag(x_2) == 8) +{ +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; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 3); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 8); +lean_dec(x_2); +x_13 = lean_expr_instantiate_rev(x_9, x_3); +lean_dec(x_9); +x_14 = lean_expr_instantiate_rev(x_10, x_3); +lean_dec(x_10); +lean_inc(x_1); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_8); +x_15 = lean_apply_6(x_1, x_8, x_14, x_4, x_5, x_6, x_7); +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); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Compiler_mkLetDecl(x_8, x_13, x_16, x_12, x_4, x_5, x_6, 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); +x_21 = lean_array_push(x_3, x_19); +x_2 = x_11; +x_3 = x_21; +x_7 = x_20; +goto _start; +} +else +{ +uint8_t x_23; +lean_dec(x_13); +lean_dec(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_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; +} +} +} +else +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_27 = lean_expr_instantiate_rev(x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_7); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp(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_Compiler_CompilerM_State_letFVars___default___closed__1; +x_8 = l_Lean_Compiler_visitLetImp_go(x_2, x_1, x_7, x_3, x_4, x_5, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_instVisitLetCompilerM___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLetImp), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_instVisitLetCompilerM() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_instVisitLetCompilerM___closed__1; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Compiler_instVisitLetReaderT___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -7017,7 +7547,7 @@ static lean_object* _init_l_Lean_Compiler_mkJump___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_Compiler_mkJump___closed__1; x_2 = l_Lean_Compiler_mkJump___closed__2; -x_3 = lean_unsigned_to_nat(234u); +x_3 = lean_unsigned_to_nat(258u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Lean_Compiler_mkJump___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7517,189 +8047,452 @@ lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___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_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda(lean_object* x_1, 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_8; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_4, x_9); +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_41 = lean_st_ref_get(x_6, x_12); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_st_ref_take(x_4, x_42); +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_is_exclusive(x_44); +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_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_47 = lean_ctor_get(x_44, 1); +lean_dec(x_47); +x_48 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +lean_ctor_set(x_44, 1, x_48); +x_49 = lean_st_ref_set(x_4, x_44, x_45); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_Lean_Compiler_visitLambdaCore(x_2, x_4, x_5, x_6, x_50); +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, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +lean_dec(x_52); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_apply_5(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_inc(x_54); +x_56 = l_Lean_Compiler_attachJp_visitLet(x_1, x_55, x_54, x_3, x_4, x_5, x_6, x_53); +if (lean_obj_tag(x_56) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -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_apply_6(x_2, x_9, x_3, x_4, x_5, x_6, x_10); -return x_11; -} -else +lean_object* x_57; lean_object* x_58; lean_object* x_59; +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); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_59 = l_Lean_Compiler_mkLetUsingScope(x_57, x_4, x_5, x_6, x_58); +if (lean_obj_tag(x_59) == 0) { -uint8_t x_12; -lean_dec(x_6); +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; uint8_t x_72; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = l_Lean_Compiler_mkLambda(x_54, x_60, x_4, x_5, x_6, x_61); lean_dec(x_5); +lean_dec(x_54); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_st_ref_get(x_6, x_64); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +x_67 = lean_st_ref_get(x_4, x_66); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_ctor_get(x_11, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_11, 1); +lean_inc(x_71); +lean_dec(x_11); +x_72 = !lean_is_exclusive(x_68); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_73 = lean_ctor_get(x_68, 1); +lean_dec(x_73); +x_74 = lean_ctor_get(x_68, 0); +lean_dec(x_74); +lean_ctor_set(x_68, 1, x_71); +lean_ctor_set(x_68, 0, x_70); +x_75 = lean_st_ref_get(x_6, x_69); +lean_dec(x_6); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_st_ref_set(x_4, x_68, x_76); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_8); -if (x_12 == 0) +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) { -return x_8; +lean_object* x_79; +x_79 = lean_ctor_get(x_77, 0); +lean_dec(x_79); +lean_ctor_set(x_77, 0, x_63); +return x_77; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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 = 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_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_63); +lean_ctor_set(x_81, 1, 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_object* x_88; lean_object* x_89; +x_82 = lean_ctor_get(x_68, 2); +lean_inc(x_82); +lean_dec(x_68); +x_83 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_83, 0, x_70); +lean_ctor_set(x_83, 1, x_71); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_st_ref_get(x_6, x_69); +lean_dec(x_6); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = lean_st_ref_set(x_4, x_83, x_85); +lean_dec(x_4); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_88 = x_86; +} else { + lean_dec_ref(x_86); + x_88 = lean_box(0); } +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_88; } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___spec__1___rarg), 7, 0); -return x_3; +lean_ctor_set(x_89, 0, x_63); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_attachJp_visitLambda___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) { -_start: +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_apply_1(x_1, x_2); -x_8 = l_Lean_Compiler_withNewScopeImp___rarg(x_7, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_90; lean_object* x_91; +lean_dec(x_54); +lean_dec(x_5); +x_90 = lean_ctor_get(x_59, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_59, 1); +lean_inc(x_91); +lean_dec(x_59); +x_13 = x_90; +x_14 = x_91; +goto block_40; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___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: +else { -lean_object* x_7; -x_7 = l_Lean_Compiler_visitLambda(x_1, x_3, x_4, x_5, x_6); -return x_7; +lean_object* x_92; lean_object* x_93; +lean_dec(x_54); +lean_dec(x_5); +x_92 = lean_ctor_get(x_56, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_56, 1); +lean_inc(x_93); +lean_dec(x_56); +x_13 = x_92; +x_14 = x_93; +goto block_40; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___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; -x_8 = lean_ctor_get(x_2, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_2, 1); -lean_inc(x_9); -lean_dec(x_2); -x_10 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +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; +x_94 = lean_ctor_get(x_44, 0); +x_95 = lean_ctor_get(x_44, 2); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_44); +x_96 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_97 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); +x_98 = lean_st_ref_set(x_4, x_97, x_45); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = l_Lean_Compiler_visitLambdaCore(x_2, x_4, x_5, x_6, x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_101, 1); +lean_inc(x_104); +lean_dec(x_101); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_11 = l_Lean_Compiler_attachJp_visitLet(x_1, x_9, x_10, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_11) == 0) +lean_inc(x_103); +x_105 = l_Lean_Compiler_attachJp_visitLet(x_1, x_104, x_103, x_3, x_4, x_5, x_6, x_102); +if (lean_obj_tag(x_105) == 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); +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_14 = l_Lean_Compiler_mkLetUsingScope(x_12, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_14) == 0) +x_108 = l_Lean_Compiler_mkLetUsingScope(x_106, x_4, x_5, x_6, x_107); +if (lean_obj_tag(x_108) == 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_Lean_Compiler_mkLambda(x_8, x_15, x_4, x_5, x_6, x_16); -lean_dec(x_6); +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; +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 = l_Lean_Compiler_mkLambda(x_103, x_109, x_4, x_5, x_6, x_110); lean_dec(x_5); +lean_dec(x_103); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_st_ref_get(x_6, x_113); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_get(x_4, x_115); +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_11, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_11, 1); +lean_inc(x_120); +lean_dec(x_11); +x_121 = lean_ctor_get(x_117, 2); +lean_inc(x_121); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + 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, 3, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_119); +lean_ctor_set(x_123, 1, x_120); +lean_ctor_set(x_123, 2, x_121); +x_124 = lean_st_ref_get(x_6, x_118); +lean_dec(x_6); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_126 = lean_st_ref_set(x_4, x_123, x_125); lean_dec(x_4); -lean_dec(x_8); -return x_17; +x_127 = lean_ctor_get(x_126, 1); +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); +} +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_112); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +else +{ +lean_object* x_130; lean_object* x_131; +lean_dec(x_103); +lean_dec(x_5); +x_130 = lean_ctor_get(x_108, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_108, 1); +lean_inc(x_131); +lean_dec(x_108); +x_13 = x_130; +x_14 = x_131; +goto block_40; +} } else { -uint8_t x_18; -lean_dec(x_8); +lean_object* x_132; lean_object* x_133; +lean_dec(x_103); +lean_dec(x_5); +x_132 = lean_ctor_get(x_105, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_105, 1); +lean_inc(x_133); +lean_dec(x_105); +x_13 = x_132; +x_14 = x_133; +goto block_40; +} +} +block_40: +{ +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_15 = lean_st_ref_get(x_6, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_4, 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_11, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_dec(x_11); +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; lean_object* x_27; uint8_t x_28; +x_23 = lean_ctor_get(x_18, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_18, 0); +lean_dec(x_24); +lean_ctor_set(x_18, 1, x_21); +lean_ctor_set(x_18, 0, x_20); +x_25 = lean_st_ref_get(x_6, x_19); lean_dec(x_6); -lean_dec(x_5); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_set(x_4, x_18, x_26); lean_dec(x_4); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -return x_14; +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 0, x_13); +return x_27; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -x_20 = lean_ctor_get(x_14, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_14); -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_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(1, 2, 0); +lean_ctor_set(x_31, 0, x_13); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -uint8_t x_22; -lean_dec(x_8); +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_32 = lean_ctor_get(x_18, 2); +lean_inc(x_32); +lean_dec(x_18); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_21); +lean_ctor_set(x_33, 2, x_32); +x_34 = lean_st_ref_get(x_6, x_19); lean_dec(x_6); -lean_dec(x_5); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_st_ref_set(x_4, x_33, x_35); lean_dec(x_4); -x_22 = !lean_is_exclusive(x_11); -if (x_22 == 0) -{ -return x_11; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_11, 0); -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_11); -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; +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); } +if (lean_is_scalar(x_38)) { + x_39 = lean_alloc_ctor(1, 2, 0); +} else { + x_39 = x_38; + lean_ctor_set_tag(x_39, 1); } +lean_ctor_set(x_39, 0, x_13); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda(lean_object* x_1, 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; -x_8 = lean_alloc_closure((void*)(l_Lean_Compiler_attachJp_visitLambda___lambda__1___boxed), 6, 1); -lean_closure_set(x_8, 0, x_2); -x_9 = lean_alloc_closure((void*)(l_Lean_Compiler_attachJp_visitLambda___lambda__2), 7, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_attachJp_visitLambda___spec__1___rarg), 7, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_attachJp_visitLambda___spec__2(x_10, x_3, x_4, x_5, x_6, x_7); -return x_11; } } LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___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) { @@ -10501,40 +11294,6 @@ lean_ctor_set(x_8, 1, x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_8); -lean_inc(x_1); -x_14 = l_Lean_Compiler_mkLetDecl(x_1, x_2, x_7, x_3, x_10, x_11, x_12, x_13); -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_15); -x_17 = lean_array_push(x_4, x_15); -x_18 = l_Lean_Compiler_isJpBinderName(x_1); -lean_dec(x_1); -if (x_18 == 0) -{ -lean_object* x_19; -lean_dec(x_15); -x_19 = l_Lean_Compiler_attachJp_visitLet(x_5, x_6, x_17, x_9, x_10, x_11, x_12, x_16); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = l_Lean_Expr_fvarId_x21(x_15); -x_21 = lean_box(0); -x_22 = l_Std_RBNode_insert___at_Lean_Compiler_attachJp_visitLet___spec__5(x_9, x_20, x_21); -x_23 = l_Lean_Compiler_attachJp_visitLet(x_5, x_6, x_17, x_22, x_10, x_11, x_12, x_16); -return x_23; -} -} -} static lean_object* _init_l_Lean_Compiler_attachJp_visitLet___closed__1() { _start: { @@ -10575,36 +11334,83 @@ lean_dec(x_11); x_16 = l_Lean_Compiler_isJpBinderName(x_9); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_box(0); -x_18 = l_Lean_Compiler_attachJp_visitLet___lambda__2(x_9, x_14, x_13, x_3, x_1, x_12, x_15, x_17, x_4, x_5, x_6, x_7, x_8); -return x_18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = l_Lean_Compiler_mkLetDecl(x_9, x_14, x_15, x_13, x_5, x_6, x_7, 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 = lean_array_push(x_3, x_18); +x_2 = x_12; +x_3 = x_20; +x_8 = x_19; +goto _start; } else { -lean_object* x_19; +lean_object* x_22; +lean_dec(x_14); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_19 = l_Lean_Compiler_attachJp_visitLambda(x_1, x_15, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_19) == 0) +x_22 = l_Lean_Compiler_attachJp_visitLambda(x_1, x_15, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_22) == 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_box(0); -x_23 = l_Lean_Compiler_attachJp_visitLet___lambda__2(x_9, x_14, x_13, x_3, x_1, x_12, x_20, x_22, x_4, x_5, x_6, x_7, x_21); -return x_23; +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_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_st_ref_get(x_7, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_5, x_26); +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); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_23); +x_31 = l_Lean_Compiler_InferType_inferType(x_23, x_30, x_6, x_7, x_29); +if (lean_obj_tag(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; +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_Lean_Compiler_mkLetDecl(x_9, x_32, x_23, x_13, x_5, x_6, x_7, x_33); +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); +lean_inc(x_35); +x_37 = lean_array_push(x_3, x_35); +x_38 = l_Lean_Expr_fvarId_x21(x_35); +x_39 = lean_box(0); +x_40 = l_Std_RBNode_insert___at_Lean_Compiler_attachJp_visitLet___spec__5(x_4, x_38, x_39); +x_2 = x_12; +x_3 = x_37; +x_4 = x_40; +x_8 = x_36; +goto _start; } else { -uint8_t x_24; -lean_dec(x_14); +uint8_t x_42; +lean_dec(x_23); lean_dec(x_12); lean_dec(x_9); lean_dec(x_7); @@ -10613,214 +11419,245 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_24 = !lean_is_exclusive(x_19); -if (x_24 == 0) +x_42 = !lean_is_exclusive(x_31); +if (x_42 == 0) { -return x_19; +return x_31; } 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_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_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_31, 0); +x_44 = lean_ctor_get(x_31, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_31); +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_12); +lean_dec(x_9); +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_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 { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_expr_instantiate_rev(x_2, x_3); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_expr_instantiate_rev(x_2, x_3); lean_dec(x_3); lean_dec(x_2); -x_29 = l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__1(x_28, x_4, x_5, x_6, x_7, x_8); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 0) +x_51 = l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__1(x_50, x_4, x_5, x_6, x_7, x_8); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_28); -x_32 = l_Lean_Compiler_isCasesApp_x3f(x_28, x_6, x_7, x_31); -if (lean_obj_tag(x_32) == 0) +lean_inc(x_50); +x_54 = l_Lean_Compiler_isCasesApp_x3f(x_50, x_6, x_7, x_53); +if (lean_obj_tag(x_54) == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -if (lean_obj_tag(x_33) == 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_34; lean_object* x_35; +lean_object* x_56; lean_object* x_57; lean_dec(x_4); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Compiler_mkJump(x_1, x_28, x_5, x_6, x_7, x_34); -return x_35; +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Lean_Compiler_mkJump(x_1, x_50, x_5, x_6, x_7, x_56); +return x_57; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_32, 1); -lean_inc(x_36); -lean_dec(x_32); -x_37 = lean_ctor_get(x_33, 0); -lean_inc(x_37); -lean_dec(x_33); -x_38 = l_Lean_Compiler_attachJp_visitCases(x_1, x_37, x_28, x_4, x_5, x_6, x_7, x_36); -return x_38; +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_54, 1); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_ctor_get(x_55, 0); +lean_inc(x_59); +lean_dec(x_55); +x_60 = l_Lean_Compiler_attachJp_visitCases(x_1, x_59, x_50, x_4, x_5, x_6, x_7, x_58); +return x_60; } } else { -uint8_t x_39; -lean_dec(x_28); +uint8_t x_61; +lean_dec(x_50); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_32); -if (x_39 == 0) +x_61 = !lean_is_exclusive(x_54); +if (x_61 == 0) { -return x_32; +return x_54; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_32, 0); -x_41 = lean_ctor_get(x_32, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_32); -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_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_54, 0); +x_63 = lean_ctor_get(x_54, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_54); +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; } } } else { -uint8_t x_43; +uint8_t x_65; lean_dec(x_1); -x_43 = !lean_is_exclusive(x_29); -if (x_43 == 0) +x_65 = !lean_is_exclusive(x_51); +if (x_65 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_ctor_get(x_29, 1); -x_45 = lean_ctor_get(x_29, 0); -lean_dec(x_45); -x_46 = lean_ctor_get(x_30, 0); -lean_inc(x_46); -lean_dec(x_30); -x_47 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_4, x_46); -lean_dec(x_46); -if (lean_obj_tag(x_47) == 0) +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_51, 1); +x_67 = lean_ctor_get(x_51, 0); +lean_dec(x_67); +x_68 = lean_ctor_get(x_52, 0); +lean_inc(x_68); +lean_dec(x_52); +x_69 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_4, x_68); +lean_dec(x_68); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -lean_free_object(x_29); -lean_dec(x_28); -x_48 = l_Lean_Compiler_attachJp_visitLet___closed__2; -x_49 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(x_48, x_4, x_5, x_6, x_7, x_44); +lean_object* x_70; lean_object* x_71; uint8_t x_72; +lean_free_object(x_51); +lean_dec(x_50); +x_70 = l_Lean_Compiler_attachJp_visitLet___closed__2; +x_71 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(x_70, x_4, x_5, x_6, x_7, x_66); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +x_72 = !lean_is_exclusive(x_71); +if (x_72 == 0) { -return x_49; +return x_71; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_49, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_49); -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; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_71, 0); +x_74 = lean_ctor_get(x_71, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_71); +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 { -lean_dec(x_47); +lean_dec(x_69); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_ctor_set(x_29, 0, x_28); -return x_29; +lean_ctor_set(x_51, 0, x_50); +return x_51; } } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_29, 1); -lean_inc(x_54); -lean_dec(x_29); -x_55 = lean_ctor_get(x_30, 0); -lean_inc(x_55); -lean_dec(x_30); -x_56 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_4, x_55); -lean_dec(x_55); -if (lean_obj_tag(x_56) == 0) +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_51, 1); +lean_inc(x_76); +lean_dec(x_51); +x_77 = lean_ctor_get(x_52, 0); +lean_inc(x_77); +lean_dec(x_52); +x_78 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_4, x_77); +lean_dec(x_77); +if (lean_obj_tag(x_78) == 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_dec(x_28); -x_57 = l_Lean_Compiler_attachJp_visitLet___closed__2; -x_58 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(x_57, x_4, x_5, x_6, x_7, x_54); +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_dec(x_50); +x_79 = l_Lean_Compiler_attachJp_visitLet___closed__2; +x_80 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(x_79, x_4, x_5, x_6, x_7, x_76); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_61 = x_58; +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_58); - x_61 = lean_box(0); + lean_dec_ref(x_80); + x_83 = lean_box(0); } -if (lean_is_scalar(x_61)) { - x_62 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(1, 2, 0); } else { - x_62 = x_61; + x_84 = x_83; } -lean_ctor_set(x_62, 0, x_59); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_82); +return x_84; } else { -lean_object* x_63; -lean_dec(x_56); +lean_object* x_85; +lean_dec(x_78); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_28); -lean_ctor_set(x_63, 1, x_54); -return x_63; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_50); +lean_ctor_set(x_85, 1, x_76); +return x_85; } } } @@ -10920,365 +11757,904 @@ if (x_30 == 0) { return x_24; } -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_24, 0); -x_32 = lean_ctor_get(x_24, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_24); -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; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_24, 0); +x_32 = lean_ctor_get(x_24, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_24); +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; +} +} +} +} +else +{ +lean_object* x_34; +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); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_6); +lean_ctor_set(x_34, 1, x_11); +return x_34; +} +} +else +{ +lean_object* x_35; +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); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_6); +lean_ctor_set(x_35, 1, x_11); +return x_35; +} +} +} +static lean_object* _init_l_Lean_Compiler_attachJp_visitCases___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.attachJp.visitCases", 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_attachJp_visitCases___closed__2() { +_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_Compiler_mkJump___closed__1; +x_2 = l_Lean_Compiler_attachJp_visitCases___closed__1; +x_3 = lean_unsigned_to_nat(317u); +x_4 = lean_unsigned_to_nat(42u); +x_5 = l_Lean_Compiler_mkJump___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_Compiler_attachJp_visitCases(lean_object* x_1, lean_object* x_2, 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; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_9); +x_11 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1; +lean_inc(x_10); +x_12 = lean_mk_array(x_10, x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_10, x_13); +lean_dec(x_10); +lean_inc(x_3); +x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_12, x_14); +x_16 = lean_st_ref_get(x_7, x_8); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_get(x_5, 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); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_22 = l_Lean_Compiler_InferType_inferType(x_1, x_21, x_6, x_7, x_20); +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) == 7) +{ +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_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 2); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Compiler_CasesInfo_updateResultingType(x_2, x_15, x_25); +lean_dec(x_25); +x_27 = lean_ctor_get(x_2, 3); +lean_inc(x_27); +lean_dec(x_2); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_27, 2); +lean_inc(x_30); +lean_dec(x_27); +lean_inc(x_28); +x_31 = l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___spec__2(x_1, x_28, x_29, x_28, x_30, x_26, x_4, x_5, x_6, x_7, x_24); +lean_dec(x_30); +lean_dec(x_28); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = l_Lean_Expr_getAppFn(x_3); +lean_dec(x_3); +x_35 = l_Lean_mkAppN(x_34, x_33); +lean_ctor_set(x_31, 0, x_35); +return x_31; +} +else +{ +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_31, 0); +x_37 = lean_ctor_get(x_31, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_31); +x_38 = l_Lean_Expr_getAppFn(x_3); +lean_dec(x_3); +x_39 = l_Lean_mkAppN(x_38, 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_37); +return x_40; +} +} +else +{ +uint8_t x_41; +lean_dec(x_3); +x_41 = !lean_is_exclusive(x_31); +if (x_41 == 0) +{ +return x_31; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_31, 0); +x_43 = lean_ctor_get(x_31, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_31); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +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_dec(x_23); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_ctor_get(x_22, 1); +lean_inc(x_45); +lean_dec(x_22); +x_46 = l_Lean_Compiler_attachJp_visitCases___closed__2; +x_47 = l_panic___at_Lean_Compiler_attachJp_visitCases___spec__1(x_46, x_4, x_5, x_6, x_7, x_45); +return x_47; +} +} +else +{ +uint8_t x_48; +lean_dec(x_15); +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_48 = !lean_is_exclusive(x_22); +if (x_48 == 0) +{ +return x_22; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_22, 0); +x_50 = lean_ctor_get(x_22, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_22); +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; +} +} } } +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___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: +{ +lean_object* x_7; +x_7 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___spec__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; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___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_34; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_7; +x_7 = l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__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); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_6); -lean_ctor_set(x_34, 1, x_11); -return x_34; +return x_7; } } -else +LEAN_EXPORT lean_object* l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_35; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); +lean_object* x_3; +x_3 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_1, x_2); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_6); -lean_ctor_set(x_35, 1, x_11); -return x_35; +return x_3; +} } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(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; } } -static lean_object* _init_l_Lean_Compiler_attachJp_visitCases___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.attachJp.visitCases", 33); -return x_1; +lean_object* x_8; +x_8 = l_Lean_Compiler_attachJp_visitLet___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; } } -static lean_object* _init_l_Lean_Compiler_attachJp_visitCases___closed__2() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___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, lean_object* x_11) { _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_Compiler_mkJump___closed__1; -x_2 = l_Lean_Compiler_attachJp_visitCases___closed__1; -x_3 = lean_unsigned_to_nat(293u); -x_4 = lean_unsigned_to_nat(42u); -x_5 = l_Lean_Compiler_mkJump___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___spec__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_5); +lean_dec(x_4); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitCases(lean_object* x_1, lean_object* x_2, lean_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_Compiler_attachJp(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_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_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_9); -x_11 = l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1; -lean_inc(x_10); -x_12 = lean_mk_array(x_10, x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_10, x_13); +lean_object* x_7; lean_object* x_8; 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_box(0); +x_8 = lean_st_ref_get(x_5, x_6); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_3, x_9); +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_inc(x_3); -x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_12, x_14); -x_16 = lean_st_ref_get(x_7, x_8); -x_17 = lean_ctor_get(x_16, 1); +x_13 = lean_st_ref_get(x_5, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_take(x_3, x_14); +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_16); -x_18 = lean_st_ref_get(x_5, 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); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_22 = l_Lean_Compiler_InferType_inferType(x_1, x_21, x_6, x_7, x_20); -if (lean_obj_tag(x_22) == 0) +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -if (lean_obj_tag(x_23) == 7) +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_16, 1); +lean_dec(x_19); +x_20 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +lean_ctor_set(x_16, 1, x_20); +x_21 = lean_st_ref_set(x_3, x_16, x_17); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_23 = l_Lean_Compiler_attachJp_visitLet(x_2, x_1, x_20, x_7, x_3, x_4, x_5, 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; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_24 = lean_ctor_get(x_22, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 2); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = l_Lean_Compiler_CasesInfo_updateResultingType(x_2, x_15, x_25); -lean_dec(x_25); -x_27 = lean_ctor_get(x_2, 3); +lean_inc(x_5); +lean_inc(x_3); +x_26 = l_Lean_Compiler_mkLetUsingScope(x_24, x_3, x_4, x_5, x_25); +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; lean_object* x_35; uint8_t x_36; +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_2); -x_28 = lean_ctor_get(x_27, 1); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_27, 2); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_5, x_28); +x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); -lean_dec(x_27); -lean_inc(x_28); -x_31 = l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___spec__2(x_1, x_28, x_29, x_28, x_30, x_26, x_4, x_5, x_6, x_7, x_24); -lean_dec(x_30); -lean_dec(x_28); -if (lean_obj_tag(x_31) == 0) +lean_dec(x_29); +x_31 = lean_st_ref_get(x_3, 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); +lean_dec(x_31); +x_34 = lean_ctor_get(x_11, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_11, 1); +lean_inc(x_35); +lean_dec(x_11); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +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_37 = lean_ctor_get(x_32, 1); +lean_dec(x_37); +x_38 = lean_ctor_get(x_32, 0); +lean_dec(x_38); +lean_ctor_set(x_32, 1, x_35); +lean_ctor_set(x_32, 0, x_34); +x_39 = lean_st_ref_get(x_5, x_33); +lean_dec(x_5); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_st_ref_set(x_3, x_32, x_40); +lean_dec(x_3); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_31, 0); -x_34 = l_Lean_Expr_getAppFn(x_3); +lean_object* x_43; +x_43 = lean_ctor_get(x_41, 0); +lean_dec(x_43); +lean_ctor_set(x_41, 0, x_27); +return x_41; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +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; +x_46 = lean_ctor_get(x_32, 2); +lean_inc(x_46); +lean_dec(x_32); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_34); +lean_ctor_set(x_47, 1, x_35); +lean_ctor_set(x_47, 2, x_46); +x_48 = lean_st_ref_get(x_5, x_33); +lean_dec(x_5); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = lean_st_ref_set(x_3, x_47, x_49); lean_dec(x_3); -x_35 = l_Lean_mkAppN(x_34, x_33); -lean_ctor_set(x_31, 0, x_35); -return x_31; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_27); +lean_ctor_set(x_53, 1, x_51); +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; uint8_t x_63; +x_54 = lean_ctor_get(x_26, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_26, 1); +lean_inc(x_55); +lean_dec(x_26); +x_56 = lean_st_ref_get(x_5, x_55); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_get(x_3, 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_11, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_11, 1); +lean_inc(x_62); +lean_dec(x_11); +x_63 = !lean_is_exclusive(x_59); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_64 = lean_ctor_get(x_59, 1); +lean_dec(x_64); +x_65 = lean_ctor_get(x_59, 0); +lean_dec(x_65); +lean_ctor_set(x_59, 1, x_62); +lean_ctor_set(x_59, 0, x_61); +x_66 = lean_st_ref_get(x_5, x_60); +lean_dec(x_5); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = lean_st_ref_set(x_3, x_59, x_67); +lean_dec(x_3); +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) +{ +lean_object* x_70; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +lean_ctor_set_tag(x_68, 1); +lean_ctor_set(x_68, 0, x_54); +return x_68; } else { -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_31, 0); -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_31); -x_38 = l_Lean_Expr_getAppFn(x_3); -lean_dec(x_3); -x_39 = l_Lean_mkAppN(x_38, 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_37); -return x_40; +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_68, 1); +lean_inc(x_71); +lean_dec(x_68); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_54); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } else { -uint8_t x_41; +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_59, 2); +lean_inc(x_73); +lean_dec(x_59); +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_61); +lean_ctor_set(x_74, 1, x_62); +lean_ctor_set(x_74, 2, x_73); +x_75 = lean_st_ref_get(x_5, x_60); +lean_dec(x_5); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_st_ref_set(x_3, x_74, x_76); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_31); -if (x_41 == 0) -{ -return x_31; +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_77); + x_79 = lean_box(0); } -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_31, 0); -x_43 = lean_ctor_get(x_31, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_31); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(1, 2, 0); +} else { + x_80 = x_79; + lean_ctor_set_tag(x_80, 1); +} +lean_ctor_set(x_80, 0, x_54); +lean_ctor_set(x_80, 1, x_78); +return x_80; } } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; +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_dec(x_4); +x_81 = lean_ctor_get(x_23, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_23, 1); +lean_inc(x_82); lean_dec(x_23); -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_45 = lean_ctor_get(x_22, 1); -lean_inc(x_45); -lean_dec(x_22); -x_46 = l_Lean_Compiler_attachJp_visitCases___closed__2; -x_47 = l_panic___at_Lean_Compiler_attachJp_visitCases___spec__1(x_46, x_4, x_5, x_6, x_7, x_45); -return x_47; -} -} -else +x_83 = lean_st_ref_get(x_5, x_82); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = lean_st_ref_get(x_3, x_84); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_88 = lean_ctor_get(x_11, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_11, 1); +lean_inc(x_89); +lean_dec(x_11); +x_90 = !lean_is_exclusive(x_86); +if (x_90 == 0) { -uint8_t x_48; -lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_6); +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_91 = lean_ctor_get(x_86, 1); +lean_dec(x_91); +x_92 = lean_ctor_get(x_86, 0); +lean_dec(x_92); +lean_ctor_set(x_86, 1, x_89); +lean_ctor_set(x_86, 0, x_88); +x_93 = lean_st_ref_get(x_5, x_87); lean_dec(x_5); -lean_dec(x_4); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +lean_dec(x_93); +x_95 = lean_st_ref_set(x_3, x_86, x_94); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_22); -if (x_48 == 0) +x_96 = !lean_is_exclusive(x_95); +if (x_96 == 0) { -return x_22; +lean_object* x_97; +x_97 = lean_ctor_get(x_95, 0); +lean_dec(x_97); +lean_ctor_set_tag(x_95, 1); +lean_ctor_set(x_95, 0, x_81); +return x_95; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_22, 0); -x_50 = lean_ctor_get(x_22, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_22); -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; -} -} +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_95, 1); +lean_inc(x_98); +lean_dec(x_95); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_81); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLambda___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: +else { -lean_object* x_7; -x_7 = l_Lean_Compiler_attachJp_visitLambda___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +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; +x_100 = lean_ctor_get(x_86, 2); +lean_inc(x_100); +lean_dec(x_86); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_88); +lean_ctor_set(x_101, 1, x_89); +lean_ctor_set(x_101, 2, x_100); +x_102 = lean_st_ref_get(x_5, x_87); lean_dec(x_5); -lean_dec(x_4); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_104 = lean_st_ref_set(x_3, x_101, x_103); lean_dec(x_3); -lean_dec(x_2); -return x_7; -} +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_106 = x_104; +} else { + lean_dec_ref(x_104); + x_106 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___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: -{ -lean_object* x_7; -x_7 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_attachJp_visitLet___spec__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; +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(1, 2, 0); +} else { + x_107 = x_106; + lean_ctor_set_tag(x_107, 1); } +lean_ctor_set(x_107, 0, x_81); +lean_ctor_set(x_107, 1, x_105); +return x_107; } -LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___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_Lean_Compiler_isJump_x3f___at_Lean_Compiler_attachJp_visitLet___spec__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_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Std_RBNode_findCore___at_Lean_Compiler_attachJp_visitLet___spec__3(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; +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_16, 0); +x_109 = lean_ctor_get(x_16, 2); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_16); +x_110 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; +x_111 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_109); +x_112 = lean_st_ref_set(x_3, x_111, x_17); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +lean_dec(x_112); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_114 = l_Lean_Compiler_attachJp_visitLet(x_2, x_1, x_110, x_7, x_3, x_4, x_5, x_113); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +lean_inc(x_5); +lean_inc(x_3); +x_117 = l_Lean_Compiler_mkLetUsingScope(x_115, x_3, x_4, x_5, x_116); +if (lean_obj_tag(x_117) == 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; 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_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = lean_st_ref_get(x_5, x_119); +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +x_122 = lean_st_ref_get(x_3, x_121); +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_11, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_11, 1); +lean_inc(x_126); +lean_dec(x_11); +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); + 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, 3, 0); +} else { + x_129 = x_128; } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___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) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_Compiler_attachJp_visitLet___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_126); +lean_ctor_set(x_129, 2, x_127); +x_130 = lean_st_ref_get(x_5, x_124); lean_dec(x_5); -lean_dec(x_4); +x_131 = lean_ctor_get(x_130, 1); +lean_inc(x_131); +lean_dec(x_130); +x_132 = lean_st_ref_set(x_3, x_129, x_131); lean_dec(x_3); -lean_dec(x_2); -return x_7; +x_133 = lean_ctor_get(x_132, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_134 = x_132; +} else { + lean_dec_ref(x_132); + 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_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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_ctor_set(x_135, 0, x_118); +lean_ctor_set(x_135, 1, x_133); +return x_135; +} +else { -lean_object* x_8; -x_8 = l_Lean_Compiler_attachJp_visitLet___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); +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; +x_136 = lean_ctor_get(x_117, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_117, 1); +lean_inc(x_137); +lean_dec(x_117); +x_138 = lean_st_ref_get(x_5, x_137); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +lean_dec(x_138); +x_140 = lean_st_ref_get(x_3, x_139); +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_140, 1); +lean_inc(x_142); +lean_dec(x_140); +x_143 = lean_ctor_get(x_11, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_11, 1); +lean_inc(x_144); +lean_dec(x_11); +x_145 = lean_ctor_get(x_141, 2); +lean_inc(x_145); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + lean_ctor_release(x_141, 2); + x_146 = x_141; +} else { + lean_dec_ref(x_141); + x_146 = lean_box(0); +} +if (lean_is_scalar(x_146)) { + x_147 = lean_alloc_ctor(0, 3, 0); +} else { + x_147 = x_146; +} +lean_ctor_set(x_147, 0, x_143); +lean_ctor_set(x_147, 1, x_144); +lean_ctor_set(x_147, 2, x_145); +x_148 = lean_st_ref_get(x_5, x_142); lean_dec(x_5); -lean_dec(x_4); +x_149 = lean_ctor_get(x_148, 1); +lean_inc(x_149); +lean_dec(x_148); +x_150 = lean_st_ref_set(x_3, x_147, x_149); lean_dec(x_3); -lean_dec(x_2); -return x_8; +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_150); + x_152 = lean_box(0); } +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; + lean_ctor_set_tag(x_153, 1); } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp_visitLet___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_Compiler_attachJp_visitLet___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); -return x_15; +lean_ctor_set(x_153, 0, x_136); +lean_ctor_set(x_153, 1, x_151); +return x_153; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___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, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_attachJp_visitCases___spec__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_5); +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_dec(x_4); -return x_12; +x_154 = lean_ctor_get(x_114, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_114, 1); +lean_inc(x_155); +lean_dec(x_114); +x_156 = lean_st_ref_get(x_5, x_155); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_158 = lean_st_ref_get(x_3, x_157); +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = lean_ctor_get(x_11, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_11, 1); +lean_inc(x_162); +lean_dec(x_11); +x_163 = lean_ctor_get(x_159, 2); +lean_inc(x_163); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + lean_ctor_release(x_159, 2); + x_164 = x_159; +} else { + lean_dec_ref(x_159); + x_164 = lean_box(0); } +if (lean_is_scalar(x_164)) { + x_165 = lean_alloc_ctor(0, 3, 0); +} else { + x_165 = x_164; } -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_attachJp___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; -x_6 = l_Lean_Compiler_withNewScopeImp___rarg(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_ctor_set(x_165, 0, x_161); +lean_ctor_set(x_165, 1, x_162); +lean_ctor_set(x_165, 2, x_163); +x_166 = lean_st_ref_get(x_5, x_160); +lean_dec(x_5); +x_167 = lean_ctor_get(x_166, 1); +lean_inc(x_167); +lean_dec(x_166); +x_168 = lean_st_ref_set(x_3, x_165, x_167); +lean_dec(x_3); +x_169 = lean_ctor_get(x_168, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_168)) { + lean_ctor_release(x_168, 0); + lean_ctor_release(x_168, 1); + x_170 = x_168; +} else { + lean_dec_ref(x_168); + x_170 = lean_box(0); } +if (lean_is_scalar(x_170)) { + x_171 = lean_alloc_ctor(1, 2, 0); +} else { + x_171 = x_170; + lean_ctor_set_tag(x_171, 1); } -static lean_object* _init_l_Lean_Compiler_attachJp___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_mkLetUsingScope), 5, 0); -return x_1; +lean_ctor_set(x_171, 0, x_154); +lean_ctor_set(x_171, 1, x_169); +return x_171; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_attachJp(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; lean_object* x_12; -x_7 = lean_box(0); -x_8 = l_Lean_Compiler_CompilerM_State_letFVars___default___closed__1; -x_9 = lean_alloc_closure((void*)(l_Lean_Compiler_attachJp_visitLet), 8, 4); -lean_closure_set(x_9, 0, x_2); -lean_closure_set(x_9, 1, x_1); -lean_closure_set(x_9, 2, x_8); -lean_closure_set(x_9, 3, x_7); -x_10 = l_Lean_Compiler_attachJp___closed__1; -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Compiler_withNewScopeImp___rarg(x_11, x_3, x_4, x_5, x_6); -return x_12; } } LEAN_EXPORT lean_object* l_Lean_Compiler_attachOptJp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -11349,6 +12725,17 @@ l_Lean_Compiler_CompilerM_State_letFVars___default = _init_l_Lean_Compiler_Compi lean_mark_persistent(l_Lean_Compiler_CompilerM_State_letFVars___default); l_Lean_Compiler_CompilerM_State_nextIdx___default = _init_l_Lean_Compiler_CompilerM_State_nextIdx___default(); lean_mark_persistent(l_Lean_Compiler_CompilerM_State_nextIdx___default); +l_Lean_Compiler_CompilerM_instInhabitedState___closed__1 = _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__1(); +lean_mark_persistent(l_Lean_Compiler_CompilerM_instInhabitedState___closed__1); +l_Lean_Compiler_CompilerM_instInhabitedState___closed__2 = _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__2(); +l_Lean_Compiler_CompilerM_instInhabitedState___closed__3 = _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__3(); +lean_mark_persistent(l_Lean_Compiler_CompilerM_instInhabitedState___closed__3); +l_Lean_Compiler_CompilerM_instInhabitedState___closed__4 = _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__4(); +lean_mark_persistent(l_Lean_Compiler_CompilerM_instInhabitedState___closed__4); +l_Lean_Compiler_CompilerM_instInhabitedState___closed__5 = _init_l_Lean_Compiler_CompilerM_instInhabitedState___closed__5(); +lean_mark_persistent(l_Lean_Compiler_CompilerM_instInhabitedState___closed__5); +l_Lean_Compiler_CompilerM_instInhabitedState = _init_l_Lean_Compiler_CompilerM_instInhabitedState(); +lean_mark_persistent(l_Lean_Compiler_CompilerM_instInhabitedState); l_Lean_Compiler_instAddMessageContextCompilerM___closed__1 = _init_l_Lean_Compiler_instAddMessageContextCompilerM___closed__1(); lean_mark_persistent(l_Lean_Compiler_instAddMessageContextCompilerM___closed__1); l_Lean_Compiler_instAddMessageContextCompilerM___closed__2 = _init_l_Lean_Compiler_instAddMessageContextCompilerM___closed__2(); @@ -11369,18 +12756,18 @@ l_Lean_Compiler_mkJpDecl___closed__2 = _init_l_Lean_Compiler_mkJpDecl___closed__ lean_mark_persistent(l_Lean_Compiler_mkJpDecl___closed__2); l_Lean_Compiler_getMaxLetVarIdx___closed__1 = _init_l_Lean_Compiler_getMaxLetVarIdx___closed__1(); lean_mark_persistent(l_Lean_Compiler_getMaxLetVarIdx___closed__1); -l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__1); -l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNames___spec__8___rarg___closed__2); -l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1 = _init_l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNames___spec__2___lambda__1___closed__1); -l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1 = _init_l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNames___spec__1___closed__1); -l_Lean_Compiler_ensureUniqueLetVarNames___closed__1 = _init_l_Lean_Compiler_ensureUniqueLetVarNames___closed__1(); -lean_mark_persistent(l_Lean_Compiler_ensureUniqueLetVarNames___closed__1); -l_Lean_Compiler_ensureUniqueLetVarNames___closed__2 = _init_l_Lean_Compiler_ensureUniqueLetVarNames___closed__2(); -lean_mark_persistent(l_Lean_Compiler_ensureUniqueLetVarNames___closed__2); +l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__1(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__1); +l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__8___rarg___closed__2); +l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1 = _init_l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Core_transform_visit___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__2___lambda__1___closed__1); +l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1 = _init_l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1___closed__1); +l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1 = _init_l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1(); +lean_mark_persistent(l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__1); +l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2 = _init_l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2(); +lean_mark_persistent(l_Lean_Compiler_ensureUniqueLetVarNamesCore___closed__2); l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__1); l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_visitMatch___spec__1___closed__2(); @@ -11439,8 +12826,6 @@ l_Lean_Compiler_attachJp_visitCases___closed__1 = _init_l_Lean_Compiler_attachJp lean_mark_persistent(l_Lean_Compiler_attachJp_visitCases___closed__1); l_Lean_Compiler_attachJp_visitCases___closed__2 = _init_l_Lean_Compiler_attachJp_visitCases___closed__2(); lean_mark_persistent(l_Lean_Compiler_attachJp_visitCases___closed__2); -l_Lean_Compiler_attachJp___closed__1 = _init_l_Lean_Compiler_attachJp___closed__1(); -lean_mark_persistent(l_Lean_Compiler_attachJp___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/Decl.c b/stage0/stdlib/Lean/Compiler/Decl.c index 084471703f46..01804b3385d5 100644 --- a/stage0/stdlib/Lean/Compiler/Decl.c +++ b/stage0/stdlib/Lean/Compiler/Decl.c @@ -64,9 +64,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_toDecl___spec__1__ LEAN_EXPORT lean_object* l_Lean_Compiler_inlineMatchers___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_inlineMatchers___closed__21; static lean_object* l_Lean_Compiler_inlineMatchers___closed__13; -lean_object* l_Lean_Compiler_check(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_inlineMatchers___lambda__1(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_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_Compiler_inlineMatchers___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Compiler_Decl_0__Lean_Compiler_normalizeAlt___spec__2(lean_object*); @@ -77,7 +78,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_Decl_0__Lean_Compiler_normali LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Compiler_Decl_0__Lean_Compiler_normalizeAlt___spec__1(lean_object*); static lean_object* l_Lean_Compiler_Decl_check___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_toLCNFType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_Decl_check___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_Decl_check___closed__1; @@ -101,6 +101,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_macroInline(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Compiler_Decl_0__Lean_Compiler_normalizeAlt___spec__1___rarg(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_Compiler_macroInline___closed__1; +lean_object* l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; static lean_object* l_Lean_Compiler_inlineMatchers___closed__10; lean_object* lean_format_pretty(lean_object*, lean_object*); @@ -128,7 +129,7 @@ lean_object* l_Lean_ConstantInfo_value_x3f(lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_inlineMatchers___closed__8; static lean_object* l_Lean_Compiler_Decl_check___closed__8; -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_macroInline___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_macroInline___lambda__1___closed__1; static lean_object* l_Lean_Compiler_inlineMatchers_inlineMatcher___closed__5; @@ -157,6 +158,7 @@ static lean_object* l_Lean_Compiler_inlineMatchers___closed__14; static lean_object* l_Lean_Compiler_Decl_checkJoinPoints___closed__1; lean_object* l_Array_ofSubarray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_inlineMatchers_inlineMatcher___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*); +static lean_object* l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2; lean_object* lean_is_unsafe_rec_name(lean_object*); static lean_object* l_Lean_Compiler_toDecl___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -164,6 +166,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_getArity(lean_object*); lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(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_Compiler_macroInline___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_toDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,7 +179,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_Decl_check___spec_ static lean_object* l_Lean_Compiler_inlineMatchers___closed__17; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); -lean_object* l_Lean_Compiler_ensureUniqueLetVarNames(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(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*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_toDecl___closed__6; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3290,7 +3293,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check(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; @@ -3562,16 +3565,6 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_check___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_2); -lean_dec(x_2); -x_7 = l_Lean_Compiler_Decl_check(x_1, x_6, x_3, x_4, x_5); -return x_7; -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_mapValue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -4065,6 +4058,22 @@ lean_dec(x_2); return x_5; } } +static lean_object* _init_l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__1___boxed), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ensureUniqueLetVarNamesCore___lambda__2___boxed), 5, 0); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_ensureUniqueLetVarNames(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -4072,134 +4081,189 @@ 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; lean_object* x_9; lean_object* x_10; +lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_6 = lean_ctor_get(x_1, 0); 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 = l_Lean_Compiler_ensureUniqueLetVarNames(x_9, x_2, x_3, x_4); -if (lean_obj_tag(x_10) == 0) +x_10 = lean_st_ref_get(x_3, x_4); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_st_mk_ref(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 = l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1; +x_17 = l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2; +lean_inc(x_3); +lean_inc(x_14); +x_18 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(x_9, x_16, x_17, x_14, x_2, x_3, x_15); +if (lean_obj_tag(x_18) == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 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 = lean_st_ref_get(x_3, x_20); +lean_dec(x_3); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_get(x_14, x_22); +lean_dec(x_14); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -lean_ctor_set(x_1, 3, x_12); -lean_ctor_set(x_10, 0, x_1); -return x_10; +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +lean_ctor_set(x_1, 3, x_19); +lean_ctor_set(x_23, 0, x_1); +return x_23; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_10, 0); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_10); -lean_ctor_set(x_1, 3, x_13); -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; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +lean_ctor_set(x_1, 3, x_19); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } else { -uint8_t x_16; +uint8_t x_28; +lean_dec(x_14); lean_free_object(x_1); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_16 = !lean_is_exclusive(x_10); -if (x_16 == 0) +lean_dec(x_3); +x_28 = !lean_is_exclusive(x_18); +if (x_28 == 0) { -return x_10; +return x_18; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_10, 0); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_10); -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_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_18, 0); +x_30 = lean_ctor_get(x_18, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_18); +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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_1, 0); -x_21 = lean_ctor_get(x_1, 1); -x_22 = lean_ctor_get(x_1, 2); -x_23 = lean_ctor_get(x_1, 3); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); +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_32 = lean_ctor_get(x_1, 0); +x_33 = lean_ctor_get(x_1, 1); +x_34 = lean_ctor_get(x_1, 2); +x_35 = lean_ctor_get(x_1, 3); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_1); -x_24 = l_Lean_Compiler_ensureUniqueLetVarNames(x_23, x_2, x_3, x_4); -if (lean_obj_tag(x_24) == 0) +x_36 = lean_st_ref_get(x_3, x_4); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_st_mk_ref(x_38, x_37); +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 = l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1; +x_43 = l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2; +lean_inc(x_3); +lean_inc(x_40); +x_44 = l_Lean_Core_transform___at_Lean_Compiler_ensureUniqueLetVarNamesCore___spec__1(x_35, x_42, x_43, x_40, x_2, x_3, x_41); +if (lean_obj_tag(x_44) == 0) { -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_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; +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 = 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_st_ref_get(x_3, x_46); +lean_dec(x_3); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_st_ref_get(x_40, x_48); +lean_dec(x_40); +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_24); - x_27 = lean_box(0); -} -x_28 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_28, 0, x_20); -lean_ctor_set(x_28, 1, x_21); -lean_ctor_set(x_28, 2, x_22); -lean_ctor_set(x_28, 3, x_25); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_49); + x_51 = lean_box(0); +} +x_52 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_52, 0, x_32); +lean_ctor_set(x_52, 1, x_33); +lean_ctor_set(x_52, 2, x_34); +lean_ctor_set(x_52, 3, x_45); +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_29 = x_27; + x_53 = x_51; } -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_50); +return x_53; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -x_30 = lean_ctor_get(x_24, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_32 = x_24; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_40); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_3); +x_54 = lean_ctor_get(x_44, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_44, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_56 = x_44; } else { - lean_dec_ref(x_24); - x_32 = lean_box(0); + lean_dec_ref(x_44); + x_56 = lean_box(0); } -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); } else { - x_33 = x_32; + x_57 = x_56; } -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_31); -return x_33; +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } } @@ -4377,6 +4441,10 @@ l_Lean_Compiler_Decl_toString___closed__1 = _init_l_Lean_Compiler_Decl_toString_ lean_mark_persistent(l_Lean_Compiler_Decl_toString___closed__1); l_Lean_Compiler_Decl_toString___closed__2 = _init_l_Lean_Compiler_Decl_toString___closed__2(); lean_mark_persistent(l_Lean_Compiler_Decl_toString___closed__2); +l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1 = _init_l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__1); +l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2 = _init_l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Decl_ensureUniqueLetVarNames___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/InferType.c b/stage0/stdlib/Lean/Compiler/InferType.c index e9cb30b9588b..ac40fa4bb73d 100644 --- a/stage0/stdlib/Lean/Compiler/InferType.c +++ b/stage0/stdlib/Lean/Compiler/InferType.c @@ -27,9 +27,10 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_i static lean_object* l_panic___at_Lean_Compiler_InferType_inferProjType___spec__3___closed__2; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___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_EXPORT lean_object* l_Lean_Compiler_InferType_inferProjType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_succ___override(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6; static lean_object* l_Lean_Compiler_InferType_getCasesResultingType_go___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___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_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___spec__4___closed__3; @@ -48,6 +49,7 @@ static lean_object* l_Lean_Compiler_InferType_getCasesResultingType_go___closed_ lean_object* l_Lean_FVarId_throwUnknown___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_InferType_getCasesResultingType_go___closed__2; +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_InferType_withLocalDeclImp___spec__2(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -63,6 +65,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inf uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___spec__4___closed__2; extern lean_object* l_Lean_levelZero; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4; static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_InferType_inferProjType___spec__1___closed__4; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -89,7 +92,7 @@ lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_InferType_inferType___closed__4; static lean_object* l_Lean_Compiler_InferType_instAddMessageContextInferTypeM___closed__5; lean_object* l_Lean_Expr_headBeta(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___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_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_InferType_instAddMessageContextInferTypeM___closed__3; static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_InferType_inferProjType___spec__1___closed__2; @@ -164,6 +167,7 @@ lean_object* l_Lean_Level_normalize(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___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_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__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_Compiler_InferType_inferProjType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_ofSubarray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_InferType_inferLambdaType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAnyType(lean_object*); lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); @@ -190,6 +194,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_instMonadInferType___rarg(lean_object*, lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___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_Lean_Compiler_instMonadInferType(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3; lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkForall(lean_object*, lean_object*, lean_object*); @@ -201,6 +206,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_InferType_inferLambdaType_go(lean_objec LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcCast___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___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* l_Lean_indentExpr(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5; static lean_object* l_panic___at_Lean_Compiler_InferType_inferProjType___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_InferType_inferFVarType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_InferType_getLevel_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -7265,37 +7271,137 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___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) { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3() { _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; uint8_t x_16; -x_10 = l_Lean_indentExpr(x_1); -x_11 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2; -x_12 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -x_13 = l_Lean_Compiler_InferType_getCasesResultingType_go___closed__4; -x_14 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_throwError___at_Lean_Compiler_InferType_inferAppType___spec__1(x_14, x_6, x_7, x_8, x_9); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" : ", 3); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4() { +_start: { -return x_15; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\nfunction type", 14); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___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: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_12 = l_Lean_Compiler_InferType_inferType(x_1, x_8, x_9, x_10, x_11); +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; 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; +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_unsigned_to_nat(0u); +x_16 = l_Array_toSubarray___rarg(x_2, x_15, x_3); +x_17 = l_Array_ofSubarray___rarg(x_16); +x_18 = l_Lean_mkAppN(x_1, x_17); +x_19 = l_Lean_indentExpr(x_18); +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2; +x_21 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4; +x_23 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_24, 0, x_4); +x_25 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6; +x_27 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_indentExpr(x_13); +x_29 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_Compiler_InferType_getCasesResultingType_go___closed__4; +x_31 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Lean_throwError___at_Lean_Compiler_InferType_inferAppType___spec__1(x_31, x_8, x_9, x_10, x_14); +lean_dec(x_10); +lean_dec(x_9); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +return x_32; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_15); -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_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(1, 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; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_12); +if (x_37 == 0) +{ +return x_12; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_12, 0); +x_39 = lean_ctor_get(x_12, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_12); +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; +} } } } @@ -7349,11 +7455,11 @@ else { lean_object* x_44; lean_object* x_45; x_44 = lean_expr_instantiate_rev_range(x_38, x_37, x_5, x_2); +lean_dec(x_38); x_45 = l_Lean_Expr_headBeta(x_44); if (lean_obj_tag(x_45) == 7) { lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_38); lean_dec(x_37); x_46 = lean_ctor_get(x_45, 2); lean_inc(x_46); @@ -7377,17 +7483,19 @@ goto block_34; else { uint8_t x_51; -lean_dec(x_45); -x_51 = l_Lean_Expr_isAnyType(x_38); +x_51 = l_Lean_Expr_isAnyType(x_45); if (x_51 == 0) { lean_object* x_52; lean_object* x_53; x_52 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); +lean_inc(x_5); +lean_inc(x_2); lean_inc(x_1); -x_53 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1(x_1, x_38, x_37, x_3, x_52, x_9, x_10, x_11, x_12); +x_53 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1(x_1, x_2, x_5, x_45, x_37, x_3, x_52, x_9, x_10, x_11, x_12); lean_dec(x_37); -lean_dec(x_38); x_18 = x_53; goto block_34; } @@ -7395,7 +7503,7 @@ else { lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_38); +lean_ctor_set(x_54, 0, x_45); lean_ctor_set(x_54, 1, x_37); x_55 = l_Array_forInUnsafe_loop___at_Lean_Compiler_InferType_inferForallType_go___spec__1___closed__1; x_56 = lean_alloc_ctor(0, 2, 0); @@ -7422,9 +7530,12 @@ if (lean_obj_tag(x_19) == 0) { uint8_t x_20; lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_20 = !lean_is_exclusive(x_18); if (x_20 == 0) @@ -7475,9 +7586,12 @@ else { uint8_t x_30; lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_30 = !lean_is_exclusive(x_18); if (x_30 == 0) @@ -7503,10 +7617,13 @@ return x_33; else { lean_object* x_59; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_59 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_59, 0, x_8); @@ -7517,10 +7634,13 @@ return x_59; else { lean_object* x_60; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_60 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_60, 0, x_8); @@ -7600,8 +7720,11 @@ lean_ctor_set(x_22, 1, x_10); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); +lean_inc(x_4); +lean_inc(x_3); lean_inc(x_2); lean_inc(x_20); +lean_inc(x_16); x_24 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2(x_9, x_16, x_21, x_20, x_10, x_20, x_14, x_23, x_2, x_3, x_4, x_19); if (lean_obj_tag(x_24) == 0) { @@ -8094,7 +8217,7 @@ static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_in 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_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___spec__4___closed__4; x_2 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___spec__4___closed__5; -x_3 = lean_unsigned_to_nat(93u); +x_3 = lean_unsigned_to_nat(94u); x_4 = lean_unsigned_to_nat(12u); x_5 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferProjType___spec__4___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9086,18 +9209,15 @@ lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_10; -x_10 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___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_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +return x_12; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___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, lean_object* x_11, lean_object* x_12) { @@ -9105,11 +9225,8 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inf { lean_object* x_13; x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__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_11); -lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); return x_13; } } @@ -9542,6 +9659,14 @@ l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lam lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__1); l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__2); +l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__3); +l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__4); +l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__5); +l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_InferType_inferAppType___spec__2___lambda__1___closed__6); l_Lean_Compiler_InferType_inferAppType___closed__1 = _init_l_Lean_Compiler_InferType_inferAppType___closed__1(); lean_mark_persistent(l_Lean_Compiler_InferType_inferAppType___closed__1); l_Lean_getConstInfo___at_Lean_Compiler_InferType_inferProjType___spec__1___closed__1 = _init_l_Lean_getConstInfo___at_Lean_Compiler_InferType_inferProjType___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/JoinPoints.c b/stage0/stdlib/Lean/Compiler/JoinPoints.c index de5faf243067..4f5c52272cd4 100644 --- a/stage0/stdlib/Lean/Compiler/JoinPoints.c +++ b/stage0/stdlib/Lean/Compiler/JoinPoints.c @@ -13,78 +13,71 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43___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_HashMapImp_moveEntries___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__48(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__22___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__5___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__15(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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*); -static lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6; -lean_object* l_Lean_Compiler_visitLetImp(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_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__49(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__7___rarg___lambda__1(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(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_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__20(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__20(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_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35___boxed(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_userName(lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__5___rarg___lambda__1(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__2___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMap_insert___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__34(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___boxed(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__16___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(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_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(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_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__16(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__18___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__16___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_jpArity___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__2(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__36(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(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_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(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* l_panic___at_Lean_LocalDecl_setBinderInfo___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__19(lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__4___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -94,45 +87,44 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__22___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__13(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(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 lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(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_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1; +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__23___rarg___lambda__1(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_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__5(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*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(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_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__4___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -lean_object* l_Lean_Compiler_withNewScopeImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__4(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2; +lean_object* l_Std_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__35(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41___boxed(lean_object*, lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__6___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__44(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__22(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,34 +132,31 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__18___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__11___rarg(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(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_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__43___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__11(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__17(lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45___boxed(lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45___boxed(lean_object*, lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__30(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__42(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -175,167 +164,171 @@ lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); lean_object* l_Lean_LocalDecl_value(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40___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_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__7(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__25(lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_inheritedTraceOptions; +static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3; lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__9(lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__17___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_Compiler_JoinPoints_forEachFVar___rarg___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(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_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_jpArity(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__23(lean_object*); -LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__19___rarg___lambda__1(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__14(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__2(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38___boxed(lean_object*, lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__2(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__2___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___boxed(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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* l_Lean_Compiler_mkLetDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__9___rarg(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__11(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__5___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32___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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__1(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32___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_panic___at_Lean_Compiler_JoinPoints_forEachFVar___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__19(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__10___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__18(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(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_Compiler_JoinPoints_forEachFVar___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31___boxed(lean_object*, lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__25(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__4___rarg___lambda__1___boxed(lean_object*, 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_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__20___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); lean_object* l_Std_mkHashMapImp___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); +static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__7___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_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__16___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__21(lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__47(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__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_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__10(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__8___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__13___rarg(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_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__43(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__24___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37___boxed(lean_object*, lean_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_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__15___rarg(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1; static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31___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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__21___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45___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_JoinPoints_forEachFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__23___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__8(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__3(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_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__6___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_getLambdaArity(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34___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_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_dbg_to_string(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__19___rarg(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__10(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_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__1(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); size_t lean_usize_modn(size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); +static lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__3___rarg(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___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_forEachFVar___spec__2___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__21(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_consumeMData(lean_object*); -LEAN_EXPORT lean_object* l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__24___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__22___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__9___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +lean_object* l_Lean_Compiler_visitMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__19___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_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__17(lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__12(lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__1(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__19___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__38(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__3___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__23(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__15(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_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -345,10 +338,11 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_ LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__31(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__6___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__1(lean_object*); @@ -357,119 +351,108 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__8___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__23___rarg(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__9(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_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__21___rarg___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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4; +LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36___boxed(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_Compiler_withNewScope___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__5(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__39___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_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__16___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__19___rarg___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35___boxed(lean_object*, 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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__20___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); -static lean_object* l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__8___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4; static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__3; static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1; static lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21___boxed(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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadCoreM; -uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__7(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__20___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__12(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__7___rarg(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__13___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__35___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__18___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__45(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__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_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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*); uint8_t l_Lean_LocalDecl_isJp(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__1___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__22(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__17___rarg(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_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__25___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__40(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__9___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__21___rarg(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_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__3(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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__5(lean_object*, size_t, size_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__11___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_Std_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__39(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_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__16(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(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_Compiler_JoinPoints_forEachFVar(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25___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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__15___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_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__9(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__17___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_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_forEachFVar___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__20___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -481,28 +464,24 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_J LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__3___rarg___lambda__1(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__9___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_forEachFVar___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_visitLetImp_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__5(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); +static lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__11___rarg___lambda__1(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(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_Compiler_JoinPoints_forEachFVar___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13___boxed(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7(lean_object*, lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__5___rarg(lean_object*, 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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__10(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__14___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__18(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19___boxed(lean_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_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__31___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___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*); @@ -513,16 +492,16 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__21___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__37(lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__2(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addDependency(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__17___rarg___lambda__1(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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__41(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -530,65 +509,64 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_ lean_object* l_Lean_Compiler_visitMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_panic___at_Lean_Compiler_JoinPoints_forEachFVar___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__14___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__14___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__10___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__29(lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__13___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__3___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__6(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__10(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_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36___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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__12___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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* l_Lean_Compiler_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(lean_object*, lean_object*, size_t, size_t, 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_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_forEachFVar___spec__3(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___lambda__11(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39___boxed(lean_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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39___boxed(lean_object*, lean_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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__4(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(lean_object*, lean_object*, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__24___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(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_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__22___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__23___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__12___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___rarg___closed__2; static lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___spec__10___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -687,7 +665,176 @@ x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg(x_1, x_2, x_3, x_4, x_7, x_5 return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___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_Lean_Compiler_JoinPoints_forEachFVar___rarg___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; 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_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_3, 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_st_ref_get(x_5, x_11); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_take(x_3, x_13); +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(x_15, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_15, 2); +lean_inc(x_18); +lean_dec(x_15); +x_19 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_20 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_18); +x_21 = lean_st_ref_set(x_3, x_20, x_16); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +lean_inc(x_5); +lean_inc(x_3); +x_23 = lean_apply_4(x_2, x_3, x_4, x_5, 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; 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_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_get(x_5, x_25); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_get(x_3, x_27); +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_10, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_10, 1); +lean_inc(x_32); +lean_dec(x_10); +x_33 = lean_ctor_get(x_29, 2); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 3, 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); +x_35 = lean_st_ref_get(x_5, x_30); +lean_dec(x_5); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_set(x_3, x_34, x_36); +lean_dec(x_3); +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_24); +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_24); +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; uint8_t x_56; +x_42 = lean_ctor_get(x_23, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_23, 1); +lean_inc(x_43); +lean_dec(x_23); +x_44 = lean_st_ref_get(x_5, x_43); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_st_ref_get(x_3, x_45); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ctor_get(x_10, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_10, 1); +lean_inc(x_50); +lean_dec(x_10); +x_51 = lean_ctor_get(x_47, 2); +lean_inc(x_51); +lean_dec(x_47); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +lean_ctor_set(x_52, 2, x_51); +x_53 = lean_st_ref_get(x_5, x_48); +lean_dec(x_5); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = lean_st_ref_set(x_3, x_52, x_54); +lean_dec(x_3); +x_56 = !lean_is_exclusive(x_55); +if (x_56 == 0) +{ +lean_object* x_57; +x_57 = lean_ctor_get(x_55, 0); +lean_dec(x_57); +lean_ctor_set_tag(x_55, 1); +lean_ctor_set(x_55, 0, x_42); +return x_55; +} +else +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_42); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___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; @@ -701,21 +848,21 @@ x_6 = lean_apply_2(x_5, lean_box(0), x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___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, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_inc(x_8); lean_inc(x_1); x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg(x_1, x_2, x_3, x_4, x_8, x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___boxed), 3, 2); +x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4___boxed), 3, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_8); x_11 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_10); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___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) { _start: { lean_object* x_7; @@ -764,7 +911,7 @@ static lean_object* _init_l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_withNewScope___rarg___lambda__1), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3), 6, 0); return x_1; } } @@ -870,7 +1017,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_28 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4___boxed), 8, 6); +x_28 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__5___boxed), 8, 6); lean_closure_set(x_28, 0, x_1); lean_closure_set(x_28, 1, x_2); lean_closure_set(x_28, 2, x_3); @@ -880,7 +1027,7 @@ lean_closure_set(x_28, 5, x_27); lean_inc(x_3); x_29 = lean_apply_2(x_3, x_7, x_28); lean_inc(x_4); -x_30 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__5), 6, 5); +x_30 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__6), 6, 5); lean_closure_set(x_30, 0, x_1); lean_closure_set(x_30, 1, x_2); lean_closure_set(x_30, 2, x_3); @@ -950,20 +1097,20 @@ lean_dec(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3(x_1, x_2, x_3); +x_4 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__4(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___rarg___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_Compiler_JoinPoints_forEachFVar___rarg___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); return x_9; } @@ -5996,1335 +6143,1360 @@ x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTai return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___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; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_10, 0, x_8); -lean_inc(x_2); -x_11 = lean_apply_2(x_2, lean_box(0), x_10); -lean_inc(x_4); -x_12 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__1___boxed), 8, 7); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_3); -lean_closure_set(x_12, 3, x_4); -lean_closure_set(x_12, 4, x_5); -lean_closure_set(x_12, 5, x_6); -lean_closure_set(x_12, 6, x_7); -x_13 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_11, x_12); -x_14 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__5; -x_15 = lean_apply_3(x_4, lean_box(0), x_14, x_13); -return x_15; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg), 8, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__1(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_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___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) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_7; lean_object* x_8; 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; uint8_t x_17; x_7 = lean_st_ref_get(x_5, x_6); x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); x_9 = lean_st_ref_get(x_3, 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; -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 = lean_local_ctx_find(x_12, x_1); -lean_ctor_set(x_9, 0, 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); +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_16 = lean_ctor_get(x_14, 0); +x_12 = lean_st_ref_get(x_5, x_11); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_take(x_3, x_13); +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_local_ctx_find(x_16, x_1); -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; -} -} -} -static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Core_instMonadCoreM; -x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1; -x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3() { -_start: +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabited___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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) { -_start: +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_15, 1); +lean_dec(x_18); +x_19 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +lean_ctor_set(x_15, 1, x_19); +x_20 = lean_st_ref_set(x_3, x_15, x_16); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +lean_inc(x_5); +lean_inc(x_3); +x_22 = lean_apply_4(x_2, x_3, x_4, x_5, x_21); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; -} -} -LEAN_EXPORT uint8_t l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(lean_object* x_1, lean_object* x_2) { -_start: +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; +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_st_ref_get(x_5, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_3, x_26); +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_10, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = !lean_is_exclusive(x_28); +if (x_32 == 0) { -if (lean_obj_tag(x_2) == 0) +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_33 = lean_ctor_get(x_28, 1); +lean_dec(x_33); +x_34 = lean_ctor_get(x_28, 0); +lean_dec(x_34); +lean_ctor_set(x_28, 1, x_31); +lean_ctor_set(x_28, 0, x_30); +x_35 = lean_st_ref_get(x_5, x_29); +lean_dec(x_5); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_set(x_3, x_28, x_36); +lean_dec(x_3); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_39; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +lean_ctor_set(x_37, 0, x_23); +return x_37; } 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 = lean_name_eq(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; +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_23); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} } else { -uint8_t x_8; -x_8 = 1; -return x_8; +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_42 = lean_ctor_get(x_28, 2); +lean_inc(x_42); +lean_dec(x_28); +x_43 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_43, 0, x_30); +lean_ctor_set(x_43, 1, x_31); +lean_ctor_set(x_43, 2, x_42); +x_44 = lean_st_ref_get(x_5, x_29); +lean_dec(x_5); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_st_ref_set(x_3, x_43, x_45); +lean_dec(x_3); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_48 = x_46; +} else { + lean_dec_ref(x_46); + x_48 = lean_box(0); } +if (lean_is_scalar(x_48)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_48; } +lean_ctor_set(x_49, 0, x_23); +lean_ctor_set(x_49, 1, x_47); +return x_49; } } -LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_2) == 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; uint8_t x_59; +x_50 = lean_ctor_get(x_22, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_22, 1); +lean_inc(x_51); +lean_dec(x_22); +x_52 = lean_st_ref_get(x_5, x_51); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_st_ref_get(x_3, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_10, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_10, 1); +lean_inc(x_58); +lean_dec(x_10); +x_59 = !lean_is_exclusive(x_55); +if (x_59 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_60 = lean_ctor_get(x_55, 1); +lean_dec(x_60); +x_61 = lean_ctor_get(x_55, 0); +lean_dec(x_61); +lean_ctor_set(x_55, 1, x_58); +lean_ctor_set(x_55, 0, x_57); +x_62 = lean_st_ref_get(x_5, x_56); +lean_dec(x_5); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = lean_st_ref_set(x_3, x_55, x_63); +lean_dec(x_3); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +lean_ctor_set_tag(x_64, 1); +lean_ctor_set(x_64, 0, x_50); +return x_64; } else { -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; 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 = lean_name_eq(x_5, x_1); -if (x_8 == 0) -{ -lean_object* x_9; -x_9 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_1, x_7); -lean_ctor_set(x_2, 2, x_9); -return x_2; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_64, 1); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_50); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} } else { -lean_free_object(x_2); -lean_dec(x_6); +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_69 = lean_ctor_get(x_55, 2); +lean_inc(x_69); +lean_dec(x_55); +x_70 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_70, 0, x_57); +lean_ctor_set(x_70, 1, x_58); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_st_ref_get(x_5, x_56); lean_dec(x_5); -return x_7; +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +lean_dec(x_71); +x_73 = lean_st_ref_set(x_3, x_70, x_72); +lean_dec(x_3); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_75 = x_73; +} else { + lean_dec_ref(x_73); + 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_tag(x_76, 1); +} +lean_ctor_set(x_76, 0, x_50); +lean_ctor_set(x_76, 1, x_74); +return x_76; } } -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_ctor_get(x_2, 0); -x_11 = lean_ctor_get(x_2, 1); -x_12 = lean_ctor_get(x_2, 2); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_2); -x_13 = lean_name_eq(x_10, x_1); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_1, x_12); -x_15 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_15, 0, x_10); -lean_ctor_set(x_15, 1, x_11); -lean_ctor_set(x_15, 2, x_14); -return x_15; } else { -lean_dec(x_11); +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_77 = lean_ctor_get(x_15, 0); +x_78 = lean_ctor_get(x_15, 2); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_15); +x_79 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_80 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_80, 2, x_78); +x_81 = lean_st_ref_set(x_3, x_80, x_16); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +lean_inc(x_5); +lean_inc(x_3); +x_83 = lean_apply_4(x_2, x_3, x_4, x_5, x_82); +if (lean_obj_tag(x_83) == 0) +{ +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_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_5, x_85); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = lean_st_ref_get(x_3, x_87); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = lean_ctor_get(x_10, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_10, 1); +lean_inc(x_92); lean_dec(x_10); -return x_12; +x_93 = lean_ctor_get(x_89, 2); +lean_inc(x_93); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + lean_ctor_release(x_89, 2); + x_94 = x_89; +} else { + lean_dec_ref(x_89); + x_94 = lean_box(0); } +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(0, 3, 0); +} else { + x_95 = x_94; } +lean_ctor_set(x_95, 0, x_91); +lean_ctor_set(x_95, 1, x_92); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_st_ref_get(x_5, x_90); +lean_dec(x_5); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = lean_st_ref_set(x_3, x_95, x_97); +lean_dec(x_3); +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_98); + x_100 = lean_box(0); } +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_100; } +lean_ctor_set(x_101, 0, x_84); +lean_ctor_set(x_101, 1, x_99); +return x_101; } -LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; size_t x_7; size_t x_8; lean_object* x_9; uint8_t x_10; -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_array_get_size(x_4); -x_6 = l_Lean_Name_hash___override(x_2); -x_7 = lean_uint64_to_usize(x_6); -x_8 = lean_usize_modn(x_7, x_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; +x_102 = lean_ctor_get(x_83, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_83, 1); +lean_inc(x_103); +lean_dec(x_83); +x_104 = lean_st_ref_get(x_5, x_103); +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_st_ref_get(x_3, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_10, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_10, 1); +lean_inc(x_110); +lean_dec(x_10); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + x_112 = x_107; +} else { + lean_dec_ref(x_107); + x_112 = lean_box(0); +} +if (lean_is_scalar(x_112)) { + x_113 = lean_alloc_ctor(0, 3, 0); +} else { + x_113 = x_112; +} +lean_ctor_set(x_113, 0, x_109); +lean_ctor_set(x_113, 1, x_110); +lean_ctor_set(x_113, 2, x_111); +x_114 = lean_st_ref_get(x_5, x_108); lean_dec(x_5); -x_9 = lean_array_uget(x_4, x_8); -x_10 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_2, x_9); -if (x_10 == 0) -{ -lean_dec(x_9); -lean_dec(x_4); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_set(x_3, x_113, x_115); lean_dec(x_3); -return x_1; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_118 = x_116; +} else { + lean_dec_ref(x_116); + x_118 = lean_box(0); } -else -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_1); -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; -x_12 = lean_ctor_get(x_1, 1); -lean_dec(x_12); -x_13 = lean_ctor_get(x_1, 0); -lean_dec(x_13); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_3, x_14); -lean_dec(x_3); -x_16 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_2, x_9); -x_17 = lean_array_uset(x_4, x_8, x_16); -lean_ctor_set(x_1, 1, x_17); -lean_ctor_set(x_1, 0, x_15); -return x_1; +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; + lean_ctor_set_tag(x_119, 1); } -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec(x_1); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_3, x_18); -lean_dec(x_3); -x_20 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_2, x_9); -x_21 = lean_array_uset(x_4, x_8, 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; +lean_ctor_set(x_119, 0, x_102); +lean_ctor_set(x_119, 1, x_117); +return x_119; } } } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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) { +static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1() { _start: { -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_apply_5(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__2), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___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; -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, 1); lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_apply_6(x_2, x_9, x_3, x_4, x_5, x_6, x_10); -return x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); +lean_closure_set(x_10, 0, x_8); +lean_inc(x_2); +x_11 = lean_apply_2(x_2, lean_box(0), x_10); +lean_inc(x_4); +x_12 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__1___boxed), 8, 7); +lean_closure_set(x_12, 0, x_1); +lean_closure_set(x_12, 1, x_2); +lean_closure_set(x_12, 2, x_3); +lean_closure_set(x_12, 3, x_4); +lean_closure_set(x_12, 4, x_5); +lean_closure_set(x_12, 5, x_6); +lean_closure_set(x_12, 6, x_7); +x_13 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_11, x_12); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1; +x_15 = lean_apply_3(x_4, lean_box(0), x_14, x_13); +return x_15; } -else -{ -uint8_t x_12; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_8); -if (x_12 == 0) -{ -return x_8; } -else +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda(lean_object* x_1) { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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 = 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_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg), 8, 0); +return x_2; } } +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___lambda__1(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_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 0); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = l_Std_mkHashSetImp___rarg(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2() { _start: { -lean_object* x_7; lean_object* x_8; -x_7 = lean_apply_1(x_1, x_2); -x_8 = l_Lean_Compiler_withNewScopeImp___rarg(x_7, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___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_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo() { _start: { -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_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(x_8, x_1, x_3, x_4, x_5, x_6, x_7); -return x_9; +lean_object* x_1; +x_1 = l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; -lean_inc(x_4); -x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(x_4, x_1, x_2, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +if (lean_obj_tag(x_2) == 0) { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else { -lean_object* x_11; -x_11 = lean_ctor_get(x_9, 0); -lean_dec(x_11); -lean_ctor_set(x_9, 0, x_4); -return x_9; +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 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; } else { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_12); -return x_13; +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; } } -else +} +} +LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_14; +lean_object* x_3; lean_object* x_4; uint64_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_array_get_size(x_3); +x_5 = l_Lean_Name_hash___override(x_2); +x_6 = lean_uint64_to_usize(x_5); +x_7 = lean_usize_modn(x_6, x_4); lean_dec(x_4); -x_14 = !lean_is_exclusive(x_9); -if (x_14 == 0) -{ +x_8 = lean_array_uget(x_3, x_7); +lean_dec(x_3); +x_9 = l_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2(x_2, x_8); +lean_dec(x_8); +lean_dec(x_2); return x_9; } +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_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_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_3, x_9); +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); +x_13 = l_Std_HashMapImp_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__1(x_12, x_1); +lean_ctor_set(x_10, 0, x_13); +return x_10; +} else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 0); -x_16 = lean_ctor_get(x_9, 1); -lean_inc(x_16); +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_dec(x_9); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); +lean_inc(x_14); +lean_dec(x_10); +x_16 = l_Std_HashMapImp_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__1(x_14, x_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); return x_17; } } } -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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_EXPORT lean_object* l_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; lean_object* x_9; -x_8 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__2___boxed), 8, 2); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_3); -x_9 = l_Lean_Compiler_visitLetImp(x_2, x_8, x_4, x_5, x_6, x_7); -return x_9; +lean_object* x_3; +x_3 = l_Std_AssocList_find_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_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) { _start: { lean_object* x_8; -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(x_2, x_1, x_3, x_4, x_5, x_6, x_7); +x_8 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f(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_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__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 uint8_t l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; -x_8 = l_Lean_Expr_consumeMData(x_1); -lean_dec(x_1); -switch (lean_obj_tag(x_8)) { -case 0: -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_8); -lean_dec(x_2); -x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_10 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_9, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -case 1: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_8, 0); -lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_apply_6(x_2, x_11, x_3, x_4, x_5, x_6, x_7); -return x_12; -} -case 2: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_8); -lean_dec(x_2); -x_13 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_14 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_13, x_3, x_4, x_5, x_6, x_7); -return x_14; +uint8_t x_3; +x_3 = 0; +return x_3; } -case 5: +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_8, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_8, 1); -lean_inc(x_16); -lean_dec(x_8); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_17 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(x_15, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_17) == 0) +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 = lean_name_eq(x_4, x_1); +if (x_6 == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_1 = x_16; -x_7 = x_18; +x_2 = x_5; goto _start; } else { -uint8_t x_20; -lean_dec(x_16); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) -{ -return x_17; +uint8_t x_8; +x_8 = 1; +return x_8; } -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; } } } -case 6: +LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_24, 0, x_8); -x_25 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_25, 0, x_24); -x_26 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__1), 7, 1); -lean_closure_set(x_26, 0, x_2); -x_27 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_27, 0, x_25); -lean_closure_set(x_27, 1, x_26); -x_28 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_27, x_3, x_4, x_5, x_6, x_7); -return x_28; -} -case 8: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_inc(x_2); -x_29 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__3), 7, 2); -lean_closure_set(x_29, 0, x_2); -lean_closure_set(x_29, 1, x_8); -x_30 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__4), 7, 1); -lean_closure_set(x_30, 0, x_2); -x_31 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_31, 0, x_29); -lean_closure_set(x_31, 1, x_30); -x_32 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_31, x_3, x_4, x_5, x_6, x_7); -return x_32; +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } -case 10: +else { -lean_object* x_33; lean_object* x_34; -lean_dec(x_8); -lean_dec(x_2); -x_33 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_34 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_33, x_3, x_4, x_5, x_6, x_7); -return x_34; -} -case 11: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) { -lean_object* x_35; -x_35 = lean_ctor_get(x_8, 2); -lean_inc(x_35); -lean_dec(x_8); -x_1 = x_35; -goto _start; +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 = lean_name_eq(x_5, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(x_1, x_7); +lean_ctor_set(x_2, 2, x_9); +return x_2; } -default: +else { -lean_object* x_37; lean_object* x_38; -lean_dec(x_8); +lean_free_object(x_2); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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_7); -return x_38; -} +return x_7; } } +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); +x_13 = lean_name_eq(x_10, x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(x_1, x_12); +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_11); +lean_ctor_set(x_15, 2, x_14); +return x_15; } -static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_private.Lean.Compiler.JoinPoints.0.Lean.Compiler.JoinPoints.JoinPointFinder.removeCandidatesContainedIn", 104); -return x_1; +lean_dec(x_11); +lean_dec(x_10); +return x_12; } } -static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2() { -_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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; -x_2 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(81u); -x_4 = lean_unsigned_to_nat(39u); -x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2; -x_11 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_10, x_2, x_3, x_4, x_5, x_9); -return x_11; -} -else +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; size_t x_7; size_t x_8; lean_object* x_9; uint8_t x_10; +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_array_get_size(x_4); +x_6 = l_Lean_Name_hash___override(x_2); +x_7 = lean_uint64_to_usize(x_6); +x_8 = lean_usize_modn(x_7, x_5); +lean_dec(x_5); +x_9 = lean_array_uget(x_4, x_8); +x_10 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2(x_2, x_9); +if (x_10 == 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; +lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_dec(x_7); -x_13 = lean_ctor_get(x_8, 0); -lean_inc(x_13); -lean_dec(x_8); -x_14 = lean_st_ref_get(x_5, x_12); -lean_dec(x_5); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_st_ref_take(x_2, x_15); -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_Lean_LocalDecl_userName(x_13); -lean_dec(x_13); -x_20 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_17, x_19); -lean_dec(x_19); -x_21 = lean_st_ref_set(x_2, x_20, x_18); -lean_dec(x_2); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_21, 0, x_24); -return x_21; -} -else -{ -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 = 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; -} -} -} -} -static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1), 6, 0); return x_1; } -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn(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: +else { -lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1; -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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: +uint8_t x_11; +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) { -lean_object* x_7; -x_7 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); +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_1, 1); +lean_dec(x_12); +x_13 = lean_ctor_get(x_1, 0); +lean_dec(x_13); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_3, x_14); lean_dec(x_3); -lean_dec(x_2); -return x_7; -} +x_16 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(x_2, x_9); +x_17 = lean_array_uset(x_4, x_8, x_16); +lean_ctor_set(x_1, 1, x_17); +lean_ctor_set(x_1, 0, x_15); +return x_1; } -LEAN_EXPORT lean_object* l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_1, x_2); -lean_dec(x_2); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_3, x_18); +lean_dec(x_3); +x_20 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(x_2, x_9); +x_21 = lean_array_uset(x_4, x_8, 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; } } -LEAN_EXPORT lean_object* l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_1, x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___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: { -lean_object* x_3; -x_3 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___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: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_9; -x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_8; lean_object* x_9; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_box(0); +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_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___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; 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 = lean_name_eq(x_4, x_1); -if (x_7 == 0) -{ -x_2 = x_6; -goto _start; -} else { -lean_object* x_9; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_dec(x_1); +lean_inc(x_6); lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; -} +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_12 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(x_10, x_2, x_3, x_4, x_5, x_6, x_7); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_1 = x_11; +x_7 = x_13; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__5(lean_object* x_1, size_t x_2, size_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_3; lean_object* x_4; uint64_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_array_get_size(x_3); -x_5 = l_Lean_Name_hash___override(x_2); -x_6 = lean_uint64_to_usize(x_5); -x_7 = lean_usize_modn(x_6, x_4); +uint8_t x_11; +x_11 = lean_usize_dec_eq(x_2, x_3); +if (x_11 == 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_dec(x_4); -x_8 = lean_array_uget(x_3, x_7); -lean_dec(x_3); -x_9 = l_Std_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__2(x_2, x_8); -lean_dec(x_8); -lean_dec(x_2); -return x_9; -} +x_12 = lean_array_uget(x_1, x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_13 = l_List_forM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__4(x_12, x_5, x_6, x_7, x_8, x_9, x_10); +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 = 1; +x_17 = lean_usize_add(x_2, x_16); +x_2 = x_17; +x_4 = x_14; +x_10 = x_15; +goto _start; } -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointFinder.findJoinPoints.goTailApp", 65); -return x_1; -} +lean_object* x_19; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_10); +return x_19; } -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2() { -_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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; -x_2 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1; -x_3 = lean_unsigned_to_nat(109u); -x_4 = lean_unsigned_to_nat(39u); -x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(lean_object* x_1, lean_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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(lean_object* x_1, 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; -x_8 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +lean_inc(x_1); +x_8 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); if (lean_obj_tag(x_9) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2; -x_12 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_11, x_3, x_4, x_5, x_6, x_10); -return x_12; +uint8_t 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); +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_8, 0, 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; uint8_t x_18; -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_8, 1); lean_inc(x_13); lean_dec(x_8); -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -lean_dec(x_9); -x_15 = lean_st_ref_get(x_6, x_13); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_st_ref_get(x_3, 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; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -x_21 = l_Lean_LocalDecl_userName(x_14); -lean_dec(x_14); -lean_inc(x_21); -x_22 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_19, x_21); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_6); -lean_dec(x_3); -x_23 = lean_box(0); -lean_ctor_set(x_17, 0, x_23); -return x_17; +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_13); +return x_15; +} } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_unsigned_to_nat(0u); -x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_25); -x_27 = lean_nat_dec_eq(x_26, x_24); -lean_dec(x_24); -lean_dec(x_26); -if (x_27 == 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; uint8_t x_25; +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_ctor_get(x_9, 0); +lean_inc(x_17); +lean_dec(x_9); +x_18 = lean_st_ref_get(x_6, x_16); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_3, 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 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1(x_21, x_1); +lean_dec(x_1); +x_24 = lean_st_ref_set(x_3, x_23, x_22); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 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; lean_object* x_34; uint8_t x_35; -lean_free_object(x_17); -x_28 = lean_st_ref_get(x_6, x_20); -lean_dec(x_6); +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_26 = lean_ctor_get(x_24, 1); +x_27 = lean_ctor_get(x_24, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_dec(x_17); x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = lean_st_ref_take(x_3, 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_30 = lean_array_get_size(x_29); +x_31 = lean_unsigned_to_nat(0u); +x_32 = lean_nat_dec_lt(x_31, x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_dec(x_30); -x_33 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_31, x_21); -lean_dec(x_21); -x_34 = lean_st_ref_set(x_3, x_33, x_32); +lean_dec(x_29); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -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 = lean_box(0); -lean_ctor_set(x_34, 0, x_37); -return x_34; +lean_dec(x_2); +x_33 = lean_box(0); +lean_ctor_set(x_24, 0, x_33); +return x_24; } 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 = lean_box(0); -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 +uint8_t x_34; +x_34 = lean_nat_dec_le(x_30, x_30); +if (x_34 == 0) { -lean_object* x_41; -lean_dec(x_21); +lean_object* x_35; +lean_dec(x_30); +lean_dec(x_29); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_41 = lean_box(0); -lean_ctor_set(x_17, 0, x_41); -return x_17; +lean_dec(x_2); +x_35 = lean_box(0); +lean_ctor_set(x_24, 0, x_35); +return x_24; +} +else +{ +size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +lean_free_object(x_24); +x_36 = 0; +x_37 = lean_usize_of_nat(x_30); +lean_dec(x_30); +x_38 = lean_box(0); +x_39 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__5(x_29, x_36, x_37, x_38, x_2, x_3, x_4, x_5, x_6, x_26); +lean_dec(x_29); +return x_39; } } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_17, 0); -x_43 = lean_ctor_get(x_17, 1); -lean_inc(x_43); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_40 = lean_ctor_get(x_24, 1); +lean_inc(x_40); +lean_dec(x_24); +x_41 = lean_ctor_get(x_17, 1); +lean_inc(x_41); lean_dec(x_17); -x_44 = l_Lean_LocalDecl_userName(x_14); -lean_dec(x_14); -lean_inc(x_44); -x_45 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_42, x_44); -if (lean_obj_tag(x_45) == 0) +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_array_get_size(x_42); +x_44 = lean_unsigned_to_nat(0u); +x_45 = lean_nat_dec_lt(x_44, x_43); +if (x_45 == 0) { lean_object* x_46; lean_object* x_47; -lean_dec(x_44); +lean_dec(x_43); +lean_dec(x_42); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); x_46 = lean_box(0); x_47 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_43); +lean_ctor_set(x_47, 1, x_40); return x_47; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_48 = lean_ctor_get(x_45, 0); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_unsigned_to_nat(0u); -x_50 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_49); -x_51 = lean_nat_dec_eq(x_50, x_48); -lean_dec(x_48); -lean_dec(x_50); -if (x_51 == 0) +uint8_t x_48; +x_48 = lean_nat_dec_le(x_43, x_43); +if (x_48 == 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; -x_52 = lean_st_ref_get(x_6, x_43); +lean_object* x_49; lean_object* x_50; +lean_dec(x_43); +lean_dec(x_42); lean_dec(x_6); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = lean_st_ref_take(x_3, x_53); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_55, x_44); -lean_dec(x_44); -x_58 = lean_st_ref_set(x_3, x_57, x_56); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; -} else { - lean_dec_ref(x_58); - x_60 = lean_box(0); -} -x_61 = lean_box(0); -if (lean_is_scalar(x_60)) { - x_62 = lean_alloc_ctor(0, 2, 0); -} else { - x_62 = x_60; -} -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_59); -return x_62; +lean_dec(x_2); +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_40); +return x_50; } else { -lean_object* x_63; lean_object* x_64; -lean_dec(x_44); -lean_dec(x_6); -lean_dec(x_3); -x_63 = lean_box(0); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_43); -return x_64; +size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; +x_51 = 0; +x_52 = lean_usize_of_nat(x_43); +lean_dec(x_43); +x_53 = lean_box(0); +x_54 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__5(x_42, x_51, x_52, x_53, x_2, x_3, x_4, x_5, x_6, x_40); +lean_dec(x_42); +return x_54; } } } } } } -LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = l_Std_AssocList_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__2(x_1, x_2); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___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_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_AssocList_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__3(x_1, x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___boxed(lean_object* x_1, lean_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_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; -x_8 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_3; +x_3 = l_Std_HashMapImp_erase___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__1(x_1, x_2); lean_dec(x_2); -return x_8; +return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___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: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__5(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +lean_dec(x_8); +x_10 = lean_st_ref_get(x_4, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_local_ctx_find(x_13, x_1); +lean_ctor_set(x_10, 0, x_14); +return x_10; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_15); +lean_dec(x_10); +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +x_18 = lean_local_ctx_find(x_17, x_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); +return x_19; } -else +} +} +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1() { +_start: { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_instMonadCoreM; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2() { +_start: { -return x_15; +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } -else +} +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; } } +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} } +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2() { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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: { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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) { +_start: +{ +if (lean_obj_tag(x_4) == 8) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_10 = lean_ctor_get(x_4, 0); lean_inc(x_10); -lean_inc(x_9); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 2); +lean_inc(x_12); +x_13 = lean_ctor_get(x_4, 3); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_4, sizeof(void*)*4 + 8); +lean_dec(x_4); +x_15 = lean_expr_instantiate_rev(x_11, x_5); +lean_dec(x_11); +x_16 = lean_expr_instantiate_rev(x_12, x_5); +lean_dec(x_12); lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; +x_17 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_16, x_1, x_2, x_3, x_6, x_7, x_8, x_9); +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; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Compiler_mkLetDecl(x_10, x_15, x_16, x_14, x_6, x_7, x_8, 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_array_push(x_5, x_20); +x_4 = x_13; +x_5 = x_22; +x_9 = x_21; goto _start; } else { -uint8_t x_21; -lean_dec(x_11); +uint8_t x_24; +lean_dec(x_16); +lean_dec(x_15); +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_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) { -return x_15; +return x_17; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_25; lean_object* x_26; lean_object* 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 = 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_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_28; lean_object* x_29; lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_28 = lean_expr_instantiate_rev(x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_9); +return x_29; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +lean_object* x_9; +x_9 = l_Lean_Expr_consumeMData(x_1); +lean_dec(x_1); +switch (lean_obj_tag(x_9)) { +case 0: { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; -goto _start; +lean_object* x_10; lean_object* x_11; +lean_dec(x_9); +lean_dec(x_2); +x_10 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_11 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; } -else +case 1: { -uint8_t x_19; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); 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_13); -if (x_19 == 0) -{ +x_13 = lean_apply_7(x_2, x_12, x_3, x_4, x_5, x_6, x_7, x_8); return x_13; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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 +case 2: { -lean_object* x_23; +lean_object* x_14; lean_object* x_15; lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; -} -} +lean_dec(x_2); +x_14 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_15 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_14, x_3, x_4, x_5, x_6, x_7, x_8); +return x_15; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5, x_6); -if (x_13 == 0) +case 5: { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_dec(x_9); +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_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_18 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_1 = x_17; +x_8 = x_19; goto _start; } else { uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_dec(x_17); +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_21 = !lean_is_exclusive(x_15); +x_21 = !lean_is_exclusive(x_18); if (x_21 == 0) { -return x_15; +return x_18; } else { lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_15); +lean_dec(x_18); x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); @@ -7332,1301 +7504,1709 @@ return x_24; } } } -else -{ -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(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) { -_start: +case 6: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +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; +x_25 = lean_st_ref_get(x_7, x_8); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_5, x_26); +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_get(x_7, x_29); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_st_ref_take(x_5, x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 2); +lean_inc(x_36); +lean_dec(x_33); +x_37 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_34); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_Compiler_visitLambda(x_9, x_5, x_6, x_7, x_40); +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_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +lean_inc(x_7); +lean_inc(x_5); +x_45 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_43); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +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; uint8_t x_60; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_st_ref_get(x_7, x_47); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = lean_st_ref_get(x_5, x_49); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_28, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_28, 1); +lean_inc(x_54); +lean_dec(x_28); +x_55 = lean_ctor_get(x_51, 2); +lean_inc(x_55); +lean_dec(x_51); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +lean_ctor_set(x_56, 2, x_55); +x_57 = lean_st_ref_get(x_7, x_52); +lean_dec(x_7); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_st_ref_set(x_5, x_56, x_58); +lean_dec(x_5); +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_object* x_61; +x_61 = lean_ctor_get(x_59, 0); +lean_dec(x_61); +lean_ctor_set(x_59, 0, x_46); +return x_59; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_46); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_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; uint8_t x_78; +x_64 = lean_ctor_get(x_45, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_45, 1); +lean_inc(x_65); +lean_dec(x_45); +x_66 = lean_st_ref_get(x_7, x_65); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = lean_st_ref_get(x_5, x_67); +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 = lean_ctor_get(x_28, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_28, 1); +lean_inc(x_72); +lean_dec(x_28); +x_73 = lean_ctor_get(x_69, 2); +lean_inc(x_73); +lean_dec(x_69); +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_73); +x_75 = lean_st_ref_get(x_7, x_70); +lean_dec(x_7); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_st_ref_set(x_5, x_74, x_76); +lean_dec(x_5); +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_77, 0); +lean_dec(x_79); +lean_ctor_set_tag(x_77, 1); +lean_ctor_set(x_77, 0, x_64); +return x_77; } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_64); +lean_ctor_set(x_81, 1, x_80); +return x_81; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(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) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +case 8: { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +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_82 = lean_st_ref_get(x_7, x_8); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +x_84 = lean_st_ref_get(x_5, x_83); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_st_ref_get(x_7, x_86); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_take(x_5, x_88); +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_ctor_get(x_90, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_90, 2); +lean_inc(x_93); +lean_dec(x_90); +x_94 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_95 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_st_ref_set(x_5, x_95, x_91); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +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_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_98 = l_Lean_Compiler_visitLetImp_go___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__5(x_2, x_3, x_4, x_9, x_94, x_5, x_6, x_7, x_97); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_object* x_99; lean_object* x_100; lean_object* 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); +lean_inc(x_7); +lean_inc(x_5); +x_101 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_99, x_2, x_3, x_4, x_5, x_6, x_7, x_100); +if (lean_obj_tag(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; 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; uint8_t x_116; +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_st_ref_get(x_7, x_103); +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_st_ref_get(x_5, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_85, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_85, 1); +lean_inc(x_110); +lean_dec(x_85); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +lean_dec(x_107); +x_112 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_110); +lean_ctor_set(x_112, 2, x_111); +x_113 = lean_st_ref_get(x_7, x_108); +lean_dec(x_7); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +lean_dec(x_113); +x_115 = lean_st_ref_set(x_5, x_112, x_114); +lean_dec(x_5); +x_116 = !lean_is_exclusive(x_115); +if (x_116 == 0) +{ +lean_object* x_117; +x_117 = lean_ctor_get(x_115, 0); +lean_dec(x_117); +lean_ctor_set(x_115, 0, x_102); +return x_115; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_115, 1); +lean_inc(x_118); +lean_dec(x_115); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_102); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +} +else { -return x_15; +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; uint8_t x_134; +x_120 = lean_ctor_get(x_101, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_101, 1); +lean_inc(x_121); +lean_dec(x_101); +x_122 = lean_st_ref_get(x_7, x_121); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = lean_st_ref_get(x_5, x_123); +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +x_127 = lean_ctor_get(x_85, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_85, 1); +lean_inc(x_128); +lean_dec(x_85); +x_129 = lean_ctor_get(x_125, 2); +lean_inc(x_129); +lean_dec(x_125); +x_130 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +lean_ctor_set(x_130, 2, x_129); +x_131 = lean_st_ref_get(x_7, x_126); +lean_dec(x_7); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +lean_dec(x_131); +x_133 = lean_st_ref_set(x_5, x_130, x_132); +lean_dec(x_5); +x_134 = !lean_is_exclusive(x_133); +if (x_134 == 0) +{ +lean_object* x_135; +x_135 = lean_ctor_get(x_133, 0); +lean_dec(x_135); +lean_ctor_set_tag(x_133, 1); +lean_ctor_set(x_133, 0, x_120); +return x_133; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_136; lean_object* x_137; +x_136 = lean_ctor_get(x_133, 1); +lean_inc(x_136); +lean_dec(x_133); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_120); +lean_ctor_set(x_137, 1, x_136); +return x_137; } } } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +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; uint8_t x_152; +lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_138 = lean_ctor_get(x_98, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_98, 1); +lean_inc(x_139); +lean_dec(x_98); +x_140 = lean_st_ref_get(x_7, x_139); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +lean_dec(x_140); +x_142 = lean_st_ref_get(x_5, x_141); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +lean_dec(x_142); +x_145 = lean_ctor_get(x_85, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_85, 1); +lean_inc(x_146); +lean_dec(x_85); +x_147 = lean_ctor_get(x_143, 2); +lean_inc(x_147); +lean_dec(x_143); +x_148 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_148, 0, x_145); +lean_ctor_set(x_148, 1, x_146); +lean_ctor_set(x_148, 2, x_147); +x_149 = lean_st_ref_get(x_7, x_144); +lean_dec(x_7); +x_150 = lean_ctor_get(x_149, 1); +lean_inc(x_150); +lean_dec(x_149); +x_151 = lean_st_ref_set(x_5, x_148, x_150); +lean_dec(x_5); +x_152 = !lean_is_exclusive(x_151); +if (x_152 == 0) +{ +lean_object* x_153; +x_153 = lean_ctor_get(x_151, 0); +lean_dec(x_153); +lean_ctor_set_tag(x_151, 1); +lean_ctor_set(x_151, 0, x_138); +return x_151; +} +else +{ +lean_object* x_154; lean_object* x_155; +x_154 = lean_ctor_get(x_151, 1); +lean_inc(x_154); +lean_dec(x_151); +x_155 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_155, 0, x_138); +lean_ctor_set(x_155, 1, x_154); +return x_155; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(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) { -_start: -{ -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +case 10: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +lean_object* x_156; lean_object* x_157; +lean_dec(x_9); +lean_dec(x_2); +x_156 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_157 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_156, x_3, x_4, x_5, x_6, x_7, x_8); +return x_157; +} +case 11: { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +lean_object* x_158; +x_158 = lean_ctor_get(x_9, 2); +lean_inc(x_158); +lean_dec(x_9); +x_1 = x_158; goto _start; } -else +default: { -uint8_t x_19; +lean_object* x_160; lean_object* x_161; 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_13); -if (x_19 == 0) -{ -return x_13; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_8); +return x_161; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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 +static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1() { +_start: { -lean_object* x_23; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_private.Lean.Compiler.JoinPoints.0.Lean.Compiler.JoinPoints.JoinPointFinder.removeCandidatesContainedIn", 104); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(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) { +static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2() { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; +x_2 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1; +x_3 = lean_unsigned_to_nat(94u); +x_4 = lean_unsigned_to_nat(39u); +x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +if (lean_obj_tag(x_9) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2; +x_12 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_11, x_2, x_3, x_4, x_5, x_6, x_10); +return x_12; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; +x_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = l_Lean_LocalDecl_userName(x_14); +lean_dec(x_14); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(x_15, x_2, x_3, x_4, x_5, x_6, x_13); +return x_16; } -else +} +} +static lean_object* _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1), 7, 0); +return x_1; +} } +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn(lean_object* x_1, 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; +x_8 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1; +x_9 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___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: { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_8; +x_8 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__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); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; -} +return x_8; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(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_EXPORT lean_object* l_Std_AssocList_foldlM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__4(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +return x_1; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; -} -else +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Name_hash___override(x_4); +x_8 = lean_uint64_to_usize(x_7); +x_9 = lean_usize_modn(x_8, x_6); +lean_dec(x_6); +x_10 = lean_array_uget(x_1, x_9); +lean_ctor_set(x_2, 2, x_10); +x_11 = lean_array_uset(x_1, x_9, x_2); +x_1 = x_11; +x_2 = x_5; +goto _start; } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* 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_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_16 = lean_array_get_size(x_1); +x_17 = l_Lean_Name_hash___override(x_13); +x_18 = lean_uint64_to_usize(x_17); +x_19 = lean_usize_modn(x_18, x_16); +lean_dec(x_16); +x_20 = lean_array_uget(x_1, x_19); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_14); +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_15; +goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(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_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; -} -else +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_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; +return x_3; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_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_AssocList_foldlM___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__4(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; } } } -else +LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_3; lean_object* x_4; 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_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_mul(x_3, x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; -} +x_6 = lean_box(0); +x_7 = lean_mk_array(x_5, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Std_HashMapImp_moveEntries___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__3(x_8, x_2, x_7); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(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_EXPORT lean_object* l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; -goto _start; +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; } else { -uint8_t x_19; -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_13); -if (x_19 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -return x_13; +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 = lean_name_eq(x_6, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +x_14 = lean_name_eq(x_11, x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(x_1, x_2, x_13); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_15); +return x_16; } else { -lean_object* x_23; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_13); +return x_17; +} } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(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_EXPORT lean_object* l_Std_HashMap_insert___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; size_t x_9; size_t x_10; lean_object* x_11; uint8_t x_12; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_array_get_size(x_6); +x_8 = l_Lean_Name_hash___override(x_2); +x_9 = lean_uint64_to_usize(x_8); +x_10 = lean_usize_modn(x_9, x_7); +x_11 = lean_array_uget(x_6, x_10); +x_12 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2(x_2, x_11); +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_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_5, x_13); +lean_dec(x_5); +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_2); +lean_ctor_set(x_15, 1, x_3); +lean_ctor_set(x_15, 2, x_11); +x_16 = lean_array_uset(x_6, x_10, x_15); +x_17 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_14); +x_18 = lean_nat_dec_le(x_17, x_7); lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_object* x_19; +lean_free_object(x_1); +x_19 = l_Std_HashMapImp_expand___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__2(x_14, x_16); +return x_19; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +lean_ctor_set(x_1, 1, x_16); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +} +else { -return x_15; +lean_object* x_20; lean_object* x_21; +lean_dec(x_7); +x_20 = l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(x_2, x_3, x_11); +x_21 = lean_array_uset(x_6, x_10, 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; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint64_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; +x_22 = lean_ctor_get(x_1, 0); +x_23 = lean_ctor_get(x_1, 1); lean_inc(x_23); lean_inc(x_22); -lean_dec(x_15); -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_1); +x_24 = lean_array_get_size(x_23); +x_25 = l_Lean_Name_hash___override(x_2); +x_26 = lean_uint64_to_usize(x_25); +x_27 = lean_usize_modn(x_26, x_24); +x_28 = lean_array_uget(x_23, x_27); +x_29 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate___spec__2(x_2, x_28); +if (x_29 == 0) +{ +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_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_22, x_30); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_2); +lean_ctor_set(x_32, 1, x_3); +lean_ctor_set(x_32, 2, x_28); +x_33 = lean_array_uset(x_23, x_27, x_32); +x_34 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_31); +x_35 = lean_nat_dec_le(x_34, x_24); +lean_dec(x_24); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = l_Std_HashMapImp_expand___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__2(x_31, x_33); +return x_36; } +else +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_31); +lean_ctor_set(x_37, 1, x_33); +return x_37; } } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_24); +x_38 = l_Std_AssocList_replace___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__5(x_2, x_3, x_28); +x_39 = lean_array_uset(x_23, x_27, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(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) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; } -else +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate(lean_object* x_1, lean_object* x_2, 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_21; +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 = l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_10, 1, x_9); +x_11 = lean_st_ref_get(x_7, x_8); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_13 = lean_st_ref_take(x_4, 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 = l_Std_HashMap_insert___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__1(x_14, x_1, x_10); +x_17 = lean_st_ref_set(x_4, x_16, x_15); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -return x_15; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___boxed(lean_object* x_1, lean_object* x_2, 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_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_9; +x_9 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate(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); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; -} +return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addDependency(lean_object* x_1, lean_object* x_2, 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_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); +lean_object* x_9; lean_object* x_10; lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_9 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_object* x_11; lean_object* x_12; +lean_dec(x_2); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_11); +return x_12; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +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_dec(x_6); +lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_st_ref_get(x_7, x_13); +lean_dec(x_7); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_take(x_4, 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_is_exclusive(x_14); +if (x_20 == 0) { -return x_15; -} -else +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_14, 1); +x_22 = l_Std_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_21, x_1); +lean_ctor_set(x_14, 1, x_22); +x_23 = l_Std_HashMap_insert___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__1(x_18, x_2, x_14); +x_24 = lean_st_ref_set(x_4, x_23, x_19); +lean_dec(x_4); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = lean_box(0); +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_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_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +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_14, 0); +x_32 = lean_ctor_get(x_14, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_14); +x_33 = l_Std_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_32, x_1); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Std_HashMap_insert___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate___spec__1(x_18, x_2, x_34); +x_36 = lean_st_ref_set(x_4, x_35, x_19); +lean_dec(x_4); +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_box(0); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__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_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); -lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +lean_dec(x_6); +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; -goto _start; +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_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_local_ctx_find(x_11, x_1); +lean_ctor_set(x_8, 0, x_12); +return x_8; } else { -uint8_t x_19; -lean_dec(x_9); +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_13); -if (x_19 == 0) -{ -return x_13; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_local_ctx_find(x_15, x_1); +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 +} +} +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1() { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_object* x_1; +x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); +return x_1; } } +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Option.get!", 11); +return x_1; } -else +} +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__3() { +_start: { -lean_object* x_23; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("value is none", 13); +return x_1; } } +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___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_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1; +x_2 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___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, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); +uint8_t x_9; +x_9 = lean_usize_dec_eq(x_3, x_4); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_array_uget(x_2, x_3); +x_11 = l_Lean_Expr_fvarId_x21(x_10); +x_12 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(x_11, x_5, x_6, x_7, x_8); +x_13 = !lean_is_exclusive(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +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); +x_16 = l_Lean_LocalDecl_userName(x_1); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4; +x_18 = l_panic___at_Lean_LocalDecl_setBinderInfo___spec__1(x_17); +x_19 = l_Lean_LocalDecl_userName(x_18); +lean_dec(x_18); +x_20 = lean_name_eq(x_19, x_16); +lean_dec(x_16); +lean_dec(x_19); +if (x_20 == 0) +{ +size_t x_21; size_t x_22; +lean_free_object(x_12); +x_21 = 1; +x_22 = lean_usize_add(x_3, x_21); +x_3 = x_22; +x_8 = x_15; goto _start; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; +uint8_t x_24; lean_object* x_25; +x_24 = 1; +x_25 = lean_box(x_24); +lean_ctor_set(x_12, 0, x_25); +return x_12; +} } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_14, 0); +lean_inc(x_26); +lean_dec(x_14); +x_27 = l_Lean_LocalDecl_userName(x_26); +lean_dec(x_26); +x_28 = lean_name_eq(x_27, x_16); +lean_dec(x_16); +lean_dec(x_27); +if (x_28 == 0) +{ +size_t x_29; size_t x_30; +lean_free_object(x_12); +x_29 = 1; +x_30 = lean_usize_add(x_3, x_29); +x_3 = x_30; +x_8 = x_15; +goto _start; } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +uint8_t x_32; lean_object* x_33; +x_32 = 1; +x_33 = lean_box(x_32); +lean_ctor_set(x_12, 0, x_33); +return x_12; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(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) { -_start: +else { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_12, 0); +x_35 = lean_ctor_get(x_12, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_12); +x_36 = l_Lean_LocalDecl_userName(x_1); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_37 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4; +x_38 = l_panic___at_Lean_LocalDecl_setBinderInfo___spec__1(x_37); +x_39 = l_Lean_LocalDecl_userName(x_38); +lean_dec(x_38); +x_40 = lean_name_eq(x_39, x_36); +lean_dec(x_36); +lean_dec(x_39); +if (x_40 == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; +size_t x_41; size_t x_42; +x_41 = 1; +x_42 = lean_usize_add(x_3, x_41); +x_3 = x_42; +x_8 = x_35; goto _start; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +uint8_t x_44; lean_object* x_45; lean_object* x_46; +x_44 = 1; +x_45 = lean_box(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_35); +return x_46; +} +} +else { -return x_15; +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = lean_ctor_get(x_34, 0); +lean_inc(x_47); +lean_dec(x_34); +x_48 = l_Lean_LocalDecl_userName(x_47); +lean_dec(x_47); +x_49 = lean_name_eq(x_48, x_36); +lean_dec(x_36); +lean_dec(x_48); +if (x_49 == 0) +{ +size_t x_50; size_t x_51; +x_50 = 1; +x_51 = lean_usize_add(x_3, x_50); +x_3 = x_51; +x_8 = x_35; +goto _start; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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; +uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_53 = 1; +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_35); +return x_55; +} } } } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +uint8_t x_56; lean_object* x_57; lean_object* x_58; +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_8); +return x_58; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 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_8, 0); +x_11 = lean_ctor_get(x_8, 1); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +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_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +uint8_t x_16; lean_object* x_17; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = 0; +x_17 = lean_box(x_16); +lean_ctor_set(x_8, 0, x_17); +return x_8; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_18; +x_18 = lean_nat_dec_le(x_13, x_13); +if (x_18 == 0) +{ +uint8_t x_19; lean_object* x_20; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; +x_19 = 0; +x_20 = lean_box(x_19); +lean_ctor_set(x_8, 0, x_20); +return x_8; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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; +size_t x_21; size_t x_22; lean_object* x_23; +lean_free_object(x_8); +x_21 = 0; +x_22 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_23 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2(x_1, x_12, x_21, x_22, x_2, x_3, x_4, x_11); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_12); +return x_23; } } } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +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_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_array_get_size(x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +if (x_29 == 0) +{ +uint8_t x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_30 = 0; +x_31 = lean_box(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_25); +return x_32; +} +else +{ +uint8_t x_33; +x_33 = lean_nat_dec_le(x_27, x_27); +if (x_33 == 0) +{ +uint8_t x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_34 = 0; +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_25); +return x_36; +} +else +{ +size_t x_37; size_t x_38; lean_object* x_39; +x_37 = 0; +x_38 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_39 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2(x_1, x_26, x_37, x_38, x_2, x_3, x_4, x_25); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_26); +return x_39; +} } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(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_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__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_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; -goto _start; +lean_object* x_6; +x_6 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__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; } -else +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___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: { -uint8_t x_19; -lean_dec(x_9); -lean_dec(x_8); +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_13); -if (x_19 == 0) -{ -return x_13; +return x_11; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_Compiler_JoinPoints_JoinPointFinder_declIsInScope___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_Compiler_JoinPoints_JoinPointFinder_declIsInScope(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; } } -else +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1() { +_start: { -lean_object* x_23; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointFinder.findJoinPoints.goTailApp", 65); +return x_1; +} } +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2() { +_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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; +x_2 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1; +x_3 = lean_unsigned_to_nat(139u); +x_4 = lean_unsigned_to_nat(39u); +x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(lean_object* x_1, lean_object* x_2, 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_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_Compiler_findDecl_x3f___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -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 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = x_17; -goto _start; +lean_dec(x_9); +x_12 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2; +x_13 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_12, x_3, x_4, x_5, x_6, x_7, x_11); +return x_13; } else { -uint8_t x_21; -lean_dec(x_11); -lean_dec(x_10); +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, 1); +lean_inc(x_14); lean_dec(x_9); -lean_dec(x_8); +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +lean_dec(x_10); +x_16 = l_Lean_LocalDecl_userName(x_15); +lean_inc(x_16); +x_17 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_findCandidate_x3f(x_16, x_3, x_4, x_5, x_6, x_7, x_14); +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_16); +lean_dec(x_15); +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_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) { -return x_15; +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_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); +x_22 = lean_ctor_get(x_17, 1); lean_inc(x_22); -lean_dec(x_15); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); +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 +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +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_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_17, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_18, 0); +lean_inc(x_28); +lean_dec(x_18); +x_29 = lean_unsigned_to_nat(0u); +x_30 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_29); +x_31 = lean_ctor_get(x_28, 0); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_nat_dec_eq(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +if (x_32 == 0) +{ +lean_object* x_33; +lean_free_object(x_17); +lean_dec(x_15); +x_33 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(x_16, x_3, x_4, x_5, x_6, x_7, x_26); +return x_33; +} +else +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_34; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_34 = lean_box(0); +lean_ctor_set(x_17, 0, x_34); +return x_17; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_free_object(x_17); +x_35 = lean_ctor_get(x_3, 0); +lean_inc(x_35); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_36 = l_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope(x_15, x_5, x_6, x_7, x_26); +lean_dec(x_15); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_unbox(x_37); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addDependency(x_16, x_35, x_3, x_4, x_5, x_6, x_7, x_39); +return x_40; } else { -lean_object* x_25; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_41; +lean_dec(x_35); +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); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_41 = !lean_is_exclusive(x_36); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_36, 0); +lean_dec(x_42); +x_43 = lean_box(0); +lean_ctor_set(x_36, 0, x_43); +return x_36; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_36, 1); +lean_inc(x_44); +lean_dec(x_36); +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_44); +return x_46; +} +} +} +} +} +else +{ +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_47 = lean_ctor_get(x_17, 1); +lean_inc(x_47); +lean_dec(x_17); +x_48 = lean_ctor_get(x_18, 0); +lean_inc(x_48); +lean_dec(x_18); +x_49 = lean_unsigned_to_nat(0u); +x_50 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_49); +x_51 = lean_ctor_get(x_48, 0); +lean_inc(x_51); +lean_dec(x_48); +x_52 = lean_nat_dec_eq(x_50, x_51); +lean_dec(x_51); +lean_dec(x_50); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_15); +x_53 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_eraseCandidate(x_16, x_3, x_4, x_5, x_6, x_7, x_47); +return x_53; +} +else +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +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_47); +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_3, 0); +lean_inc(x_56); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_57 = l_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope(x_15, x_5, x_6, x_7, x_47); +lean_dec(x_15); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_unbox(x_58); +lean_dec(x_58); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +lean_dec(x_57); +x_61 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addDependency(x_16, x_56, x_3, x_4, x_5, x_6, x_7, x_60); +return x_61; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_56); +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_62 = lean_ctor_get(x_57, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_63 = x_57; +} else { + lean_dec_ref(x_57); + x_63 = lean_box(0); +} +x_64 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +} } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(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_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___boxed(lean_object* x_1, lean_object* x_2, 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_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +lean_object* x_9; +x_9 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(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) { +_start: { -lean_object* x_14; lean_object* x_15; +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -8634,26 +9214,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8661,29 +9242,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8691,23 +9273,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -8715,26 +9298,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8742,29 +9326,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8772,97 +9357,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -8870,26 +9459,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8897,29 +9487,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8927,23 +9518,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -8951,26 +9543,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -8978,29 +9571,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9008,23 +9602,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9032,26 +9627,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9059,29 +9655,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9089,97 +9686,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9187,26 +9788,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9214,29 +9816,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9244,23 +9847,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9268,26 +9872,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9295,29 +9900,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9325,23 +9931,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9349,26 +9956,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9376,29 +9984,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9406,97 +10015,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9504,26 +10117,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9531,29 +10145,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9561,23 +10176,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9585,26 +10201,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9612,29 +10229,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9642,23 +10260,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9666,26 +10285,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9693,29 +10313,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9723,97 +10344,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9821,26 +10446,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9848,29 +10474,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9878,23 +10505,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9902,26 +10530,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9929,29 +10558,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9959,23 +10589,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -9983,26 +10614,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10010,29 +10642,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10040,97 +10673,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10138,26 +10775,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10165,29 +10803,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10195,23 +10834,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10219,26 +10859,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10246,29 +10887,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10276,23 +10918,24 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10300,26 +10943,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10327,29 +10971,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10357,97 +11002,101 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(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: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_dec(x_5); -x_12 = lean_array_uget(x_2, x_3); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_13 = lean_apply_6(x_1, x_12, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_5 = x_14; -x_10 = x_15; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; goto _start; } else { -uint8_t x_19; +uint8_t x_20; +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_13); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -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_dec(x_14); +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_23; +lean_object* x_24; +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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_5); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_5, x_6); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -x_14 = lean_array_uget(x_4, x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10455,26 +11104,27 @@ lean_inc(x_8); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_5, x_18); -x_5 = x_19; -x_7 = x_16; -x_12 = 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10482,29 +11132,30 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_15; +return x_16; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); +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_inc(x_22); -lean_dec(x_15); -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_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_25; +lean_object* x_26; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10512,1722 +11163,1712 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(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) { _start: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_7); +x_15 = lean_array_uget(x_4, x_5); lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -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); -lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); +uint8_t x_22; lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +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_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_26; +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else -{ -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(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) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_7); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); +uint8_t x_22; lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) -{ -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +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_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_26; +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_box(0); -lean_ctor_set(x_14, 0, x_35); -return x_14; -} -else -{ -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__3(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(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: { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); +lean_inc(x_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); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -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; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) -{ -lean_dec(x_47); -lean_dec(x_13); +uint8_t x_20; +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); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_24; +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); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(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) { +_start: { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; -} -else +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; -} -} -} +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_60; -lean_dec(x_13); +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_43; +return x_16; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} -} +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_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); +lean_object* x_26; lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(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) { +_start: { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; -} -else +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; -} -} +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); +uint8_t x_22; lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; -} -else -{ -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -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_64); -return x_86; +return x_16; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__3(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +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 { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); +lean_object* x_26; lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) -{ -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(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) { +_start: { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; -} -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; -} -else +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; -} -} +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} +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 { -uint8_t x_108; -lean_dec(x_13); +lean_object* x_26; 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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) -{ -return x_14; -} -else -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(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_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_ctor_get(x_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_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_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_14) == 0) { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -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; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) -{ -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +uint8_t x_20; +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); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ return x_14; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) -{ -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_24; +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); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(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) { +_start: { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_box(0); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +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 { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); +lean_object* x_26; lean_dec(x_12); -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; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) -{ -lean_dec(x_47); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(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) { +_start: { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +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_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_26; +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(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) { +_start: { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; -} -else +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; -} -} -} +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_60; -lean_dec(x_13); +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_43; +return x_16; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} -} +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_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); +lean_object* x_26; lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(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: { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +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); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_24; +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); lean_dec(x_1); -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_64); -return x_83; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(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) { +_start: { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +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 { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); +lean_object* x_26; lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) -{ -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(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) { +_start: { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; -} -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +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_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_26; +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(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) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_7); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_108; -lean_dec(x_13); +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_14; +return x_16; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +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; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object* x_7, lean_object* x_8, lean_object* x_9) { -_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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); +lean_object* x_26; +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(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: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); lean_inc(x_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_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_14) == 0) { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -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; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) -{ -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +uint8_t x_20; +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); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ return x_14; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) -{ -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_24; +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); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(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) { +_start: { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_box(0); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +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 { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); +lean_object* x_26; lean_dec(x_12); -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; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) -{ -lean_dec(x_47); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(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) { +_start: { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +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_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_26; +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(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) { +_start: { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; -} -else +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; -} -} -} +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; } else { -uint8_t x_60; -lean_dec(x_13); +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_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_43; +return x_16; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} -} +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_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); +lean_object* x_26; lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(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: { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_14 = lean_apply_7(x_1, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_3 = x_18; +x_5 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +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); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_24; +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); lean_dec(x_1); -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_64); -return x_83; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(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) { +_start: { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_5, x_6); +if (x_14 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_15 = lean_array_uget(x_4, x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_15); +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_5, x_19); +x_5 = x_20; +x_7 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +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_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +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 { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); +lean_object* x_26; lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) -{ -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +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_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_13); +return x_26; } -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +if (lean_obj_tag(x_4) == 8) { -lean_object* x_101; -lean_dec(x_97); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 2); +lean_inc(x_12); +x_13 = lean_ctor_get(x_4, 3); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_4, sizeof(void*)*4 + 8); +lean_dec(x_4); +x_15 = lean_expr_instantiate_rev(x_11, x_5); +lean_dec(x_11); +x_16 = lean_expr_instantiate_rev(x_12, x_5); +lean_dec(x_12); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_10); +x_17 = lean_apply_8(x_1, x_10, x_16, x_2, x_3, x_6, x_7, x_8, x_9); +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); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Compiler_mkLetDecl(x_10, x_15, x_18, x_14, 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_array_push(x_5, x_21); +x_4 = x_13; +x_5 = x_23; +x_9 = x_22; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_15); lean_dec(x_13); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -12235,380 +12876,591 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; -} -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +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_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_29; lean_object* x_30; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} -} +x_29 = lean_expr_instantiate_rev(x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_9); +return x_30; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_108; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_11; +x_11 = l_Lean_Expr_consumeMData(x_4); +switch (lean_obj_tag(x_11)) { +case 0: +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_11); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) -{ -return x_14; +x_12 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_13 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_12, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } -else +case 2: { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} -} -} +lean_object* x_14; lean_object* x_15; +lean_dec(x_11); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_15 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_14, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +case 5: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); +lean_object* x_16; +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +switch (lean_obj_tag(x_16)) { +case 0: +{ +lean_object* x_17; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +x_17 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +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) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_3); +lean_dec(x_1); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_19); +return x_20; +} +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_47; lean_object* x_48; 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; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_st_ref_get(x_9, x_21); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_get(x_7, x_24); +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_66 = lean_st_ref_get(x_9, x_27); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = lean_st_ref_take(x_7, x_67); +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 = 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 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_72); +x_75 = lean_st_ref_set(x_7, x_74, x_70); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_Lean_Compiler_visitMatch(x_11, x_22, x_7, x_8, x_9, x_76); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_ctor_get(x_78, 0); +lean_inc(x_81); +lean_dec(x_78); +x_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_79, 1); +lean_inc(x_83); +lean_dec(x_79); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_84 = lean_apply_7(x_2, x_81, x_5, x_6, x_7, x_8, x_9, x_80); +if (lean_obj_tag(x_84) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = lean_array_get_size(x_82); +x_87 = lean_unsigned_to_nat(0u); +x_88 = lean_nat_dec_lt(x_87, x_86); +if (x_88 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_89; uint8_t x_90; +lean_dec(x_86); +lean_dec(x_82); +x_89 = lean_array_get_size(x_83); +x_90 = lean_nat_dec_lt(x_87, x_89); +if (x_90 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_91; +lean_dec(x_89); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_91 = lean_box(0); +x_28 = x_91; +x_29 = x_85; +goto block_46; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_92; +x_92 = lean_nat_dec_le(x_89, x_89); +if (x_92 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_93; +lean_dec(x_89); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_93 = lean_box(0); +x_28 = x_93; +x_29 = x_85; +goto block_46; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_94; size_t x_95; lean_object* x_96; lean_object* x_97; +x_94 = 0; +x_95 = lean_usize_of_nat(x_89); +lean_dec(x_89); +x_96 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_97 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_83, x_94, x_95, x_96, x_5, x_6, x_7, x_8, x_9, x_85); +lean_dec(x_83); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_28 = x_98; +x_29 = x_99; +goto block_46; +} +else +{ +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_97, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_97, 1); +lean_inc(x_101); +lean_dec(x_97); +x_47 = x_100; +x_48 = x_101; +goto block_65; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_102; +x_102 = lean_nat_dec_le(x_86, x_86); +if (x_102 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_103; uint8_t x_104; +lean_dec(x_86); +lean_dec(x_82); +x_103 = lean_array_get_size(x_83); +x_104 = lean_nat_dec_lt(x_87, x_103); +if (x_104 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_105; +lean_dec(x_103); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_105 = lean_box(0); +x_28 = x_105; +x_29 = x_85; +goto block_46; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_106; +x_106 = lean_nat_dec_le(x_103, x_103); +if (x_106 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_107; +lean_dec(x_103); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_107 = lean_box(0); +x_28 = x_107; +x_29 = x_85; +goto block_46; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_108; size_t x_109; lean_object* x_110; lean_object* x_111; +x_108 = 0; +x_109 = lean_usize_of_nat(x_103); +lean_dec(x_103); +x_110 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_111 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__3(x_1, x_2, x_3, x_83, x_108, x_109, x_110, x_5, x_6, x_7, x_8, x_9, x_85); +lean_dec(x_83); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_28 = x_112; +x_29 = x_113; +goto block_46; +} +else +{ +lean_object* x_114; lean_object* x_115; +x_114 = lean_ctor_get(x_111, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_111, 1); +lean_inc(x_115); +lean_dec(x_111); +x_47 = x_114; +x_48 = x_115; +goto block_65; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_116; size_t x_117; lean_object* x_118; lean_object* x_119; +x_116 = 0; +x_117 = lean_usize_of_nat(x_86); +lean_dec(x_86); +x_118 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_119 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(x_2, x_82, x_116, x_117, x_118, x_5, x_6, x_7, x_8, x_9, x_85); +lean_dec(x_82); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_120; lean_object* x_121; uint8_t x_122; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = lean_array_get_size(x_83); +x_122 = lean_nat_dec_lt(x_87, x_121); +if (x_122 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_121); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_28 = x_118; +x_29 = x_120; +goto block_46; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_123; +x_123 = lean_nat_dec_le(x_121, x_121); +if (x_123 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_121); +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_28 = x_118; +x_29 = x_120; +goto block_46; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_124; lean_object* x_125; +x_124 = lean_usize_of_nat(x_121); +lean_dec(x_121); +lean_inc(x_9); +lean_inc(x_7); +x_125 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_1, x_2, x_3, x_83, x_116, x_124, x_118, x_5, x_6, x_7, x_8, x_9, x_120); +lean_dec(x_83); +if (lean_obj_tag(x_125) == 0) +{ +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_28 = x_126; +x_29 = x_127; +goto block_46; +} +else +{ +lean_object* x_128; lean_object* x_129; +x_128 = lean_ctor_get(x_125, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_125, 1); +lean_inc(x_129); +lean_dec(x_125); +x_47 = x_128; +x_48 = x_129; +goto block_65; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_130; lean_object* x_131; +lean_dec(x_83); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_130 = lean_ctor_get(x_119, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_119, 1); +lean_inc(x_131); +lean_dec(x_119); +x_47 = x_130; +x_48 = x_131; +goto block_65; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_132; lean_object* x_133; +lean_dec(x_83); +lean_dec(x_82); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_132 = lean_ctor_get(x_84, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_84, 1); +lean_inc(x_133); +lean_dec(x_84); +x_47 = x_132; +x_48 = x_133; +goto block_65; } -else +block_46: { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; -} -} -} -} -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; uint8_t x_42; +x_30 = lean_st_ref_get(x_9, x_29); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_st_ref_get(x_7, x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_ctor_get(x_26, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_26, 1); +lean_inc(x_36); +lean_dec(x_26); +x_37 = lean_ctor_get(x_33, 2); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 2, x_37); +x_39 = lean_st_ref_get(x_9, x_34); +lean_dec(x_9); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_st_ref_set(x_7, x_38, x_40); +lean_dec(x_7); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -uint8_t x_60; -lean_dec(x_13); +lean_object* x_43; +x_43 = lean_ctor_get(x_41, 0); +lean_dec(x_43); +lean_ctor_set(x_41, 0, x_28); +return x_41; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_28); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +block_65: +{ +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; uint8_t x_61; +x_49 = lean_st_ref_get(x_9, x_48); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = lean_st_ref_get(x_7, x_50); +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_26, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_26, 1); +lean_inc(x_55); +lean_dec(x_26); +x_56 = lean_ctor_get(x_52, 2); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +lean_ctor_set(x_57, 2, x_56); +x_58 = lean_st_ref_get(x_9, x_53); +lean_dec(x_9); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_st_ref_set(x_7, x_57, x_59); +lean_dec(x_7); +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_60, 0); +lean_dec(x_62); +lean_ctor_set_tag(x_60, 1); +lean_ctor_set(x_60, 0, x_47); +return x_60; +} +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_47); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +} +else +{ +uint8_t x_134; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -12616,289 +13468,592 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_134 = !lean_is_exclusive(x_17); +if (x_134 == 0) { -return x_43; +return x_17; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_17, 0); +x_136 = lean_ctor_get(x_17, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_17); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +return x_137; } } } +case 1: +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_dec(x_3); +x_138 = lean_ctor_get(x_11, 1); +lean_inc(x_138); +x_139 = lean_ctor_get(x_16, 0); +lean_inc(x_139); +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_140 = lean_apply_8(x_1, x_139, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_140) == 0) +{ +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +lean_dec(x_140); +x_142 = lean_apply_7(x_2, x_138, x_5, x_6, x_7, x_8, x_9, x_141); +return x_142; } +else +{ +uint8_t x_143; +lean_dec(x_138); +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_143 = !lean_is_exclusive(x_140); +if (x_143 == 0) +{ +return x_140; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_140, 0); +x_145 = lean_ctor_get(x_140, 1); +lean_inc(x_145); +lean_inc(x_144); +lean_dec(x_140); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +return x_146; +} +} +} +case 2: { -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +lean_object* x_147; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_147 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_147) == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_148; +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; +lean_dec(x_3); +lean_dec(x_1); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_149); +return x_150; +} +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_177; lean_object* x_178; 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; +x_151 = lean_ctor_get(x_147, 1); +lean_inc(x_151); +lean_dec(x_147); +x_152 = lean_ctor_get(x_148, 0); +lean_inc(x_152); +lean_dec(x_148); +x_153 = lean_st_ref_get(x_9, x_151); +x_154 = lean_ctor_get(x_153, 1); +lean_inc(x_154); +lean_dec(x_153); +x_155 = lean_st_ref_get(x_7, x_154); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_196 = lean_st_ref_get(x_9, x_157); +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +x_198 = lean_st_ref_take(x_7, x_197); +x_199 = lean_ctor_get(x_198, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_198, 1); +lean_inc(x_200); +lean_dec(x_198); +x_201 = lean_ctor_get(x_199, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_199, 2); +lean_inc(x_202); +lean_dec(x_199); +x_203 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_204 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_204, 0, x_201); +lean_ctor_set(x_204, 1, x_203); +lean_ctor_set(x_204, 2, x_202); +x_205 = lean_st_ref_set(x_7, x_204, x_200); +x_206 = lean_ctor_get(x_205, 1); +lean_inc(x_206); +lean_dec(x_205); +x_207 = l_Lean_Compiler_visitMatch(x_11, x_152, x_7, x_8, x_9, x_206); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_208, 1); +lean_inc(x_209); +x_210 = lean_ctor_get(x_207, 1); +lean_inc(x_210); +lean_dec(x_207); +x_211 = lean_ctor_get(x_208, 0); +lean_inc(x_211); +lean_dec(x_208); +x_212 = lean_ctor_get(x_209, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_209, 1); +lean_inc(x_213); +lean_dec(x_209); +lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_214 = lean_apply_7(x_2, x_211, x_5, x_6, x_7, x_8, x_9, x_210); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; +x_215 = lean_ctor_get(x_214, 1); +lean_inc(x_215); +lean_dec(x_214); +x_216 = lean_array_get_size(x_212); +x_217 = lean_unsigned_to_nat(0u); +x_218 = lean_nat_dec_lt(x_217, x_216); +if (x_218 == 0) +{ +lean_object* x_219; uint8_t x_220; +lean_dec(x_216); +lean_dec(x_212); +x_219 = lean_array_get_size(x_213); +x_220 = lean_nat_dec_lt(x_217, x_219); +if (x_220 == 0) +{ +lean_object* x_221; +lean_dec(x_219); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_221 = lean_box(0); +x_158 = x_221; +x_159 = x_215; +goto block_176; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_222; +x_222 = lean_nat_dec_le(x_219, x_219); +if (x_222 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_223; +lean_dec(x_219); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_223 = lean_box(0); +x_158 = x_223; +x_159 = x_215; +goto block_176; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +size_t x_224; size_t x_225; lean_object* x_226; lean_object* x_227; +x_224 = 0; +x_225 = lean_usize_of_nat(x_219); +lean_dec(x_219); +x_226 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_227 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(x_1, x_2, x_3, x_213, x_224, x_225, x_226, x_5, x_6, x_7, x_8, x_9, x_215); +lean_dec(x_213); +if (lean_obj_tag(x_227) == 0) +{ +lean_object* x_228; lean_object* x_229; +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_227, 1); +lean_inc(x_229); +lean_dec(x_227); +x_158 = x_228; +x_159 = x_229; +goto block_176; +} +else +{ +lean_object* x_230; lean_object* x_231; +x_230 = lean_ctor_get(x_227, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_227, 1); +lean_inc(x_231); +lean_dec(x_227); +x_177 = x_230; +x_178 = x_231; +goto block_195; +} } } } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +uint8_t x_232; +x_232 = lean_nat_dec_le(x_216, x_216); +if (x_232 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +lean_object* x_233; uint8_t x_234; +lean_dec(x_216); +lean_dec(x_212); +x_233 = lean_array_get_size(x_213); +x_234 = lean_nat_dec_lt(x_217, x_233); +if (x_234 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_235; +lean_dec(x_233); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_235 = lean_box(0); +x_158 = x_235; +x_159 = x_215; +goto block_176; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_236; +x_236 = lean_nat_dec_le(x_233, x_233); +if (x_236 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_237; +lean_dec(x_233); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_237 = lean_box(0); +x_158 = x_237; +x_159 = x_215; +goto block_176; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +size_t x_238; size_t x_239; lean_object* x_240; lean_object* x_241; +x_238 = 0; +x_239 = lean_usize_of_nat(x_233); +lean_dec(x_233); +x_240 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_241 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(x_1, x_2, x_3, x_213, x_238, x_239, x_240, x_5, x_6, x_7, x_8, x_9, x_215); +lean_dec(x_213); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; lean_object* x_243; +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_158 = x_242; +x_159 = x_243; +goto block_176; +} +else +{ +lean_object* x_244; lean_object* x_245; +x_244 = lean_ctor_get(x_241, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_241, 1); +lean_inc(x_245); +lean_dec(x_241); +x_177 = x_244; +x_178 = x_245; +goto block_195; +} } } } else { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); +size_t x_246; size_t x_247; lean_object* x_248; lean_object* x_249; +x_246 = 0; +x_247 = lean_usize_of_nat(x_216); +lean_dec(x_216); +x_248 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +lean_inc(x_2); +x_249 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(x_2, x_212, x_246, x_247, x_248, x_5, x_6, x_7, x_8, x_9, x_215); +lean_dec(x_212); +if (lean_obj_tag(x_249) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_250; lean_object* x_251; uint8_t x_252; +x_250 = lean_ctor_get(x_249, 1); +lean_inc(x_250); +lean_dec(x_249); +x_251 = lean_array_get_size(x_213); +x_252 = lean_nat_dec_lt(x_217, x_251); +if (x_252 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_251); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_158 = x_248; +x_159 = x_250; +goto block_176; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_253; +x_253 = lean_nat_dec_le(x_251, x_251); +if (x_253 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_251); +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +x_158 = x_248; +x_159 = x_250; +goto block_176; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +else +{ +size_t x_254; lean_object* x_255; +x_254 = lean_usize_of_nat(x_251); +lean_dec(x_251); +lean_inc(x_9); +lean_inc(x_7); +x_255 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_1, x_2, x_3, x_213, x_246, x_254, x_248, x_5, x_6, x_7, x_8, x_9, x_250); +lean_dec(x_213); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_255, 1); +lean_inc(x_257); +lean_dec(x_255); +x_158 = x_256; +x_159 = x_257; +goto block_176; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_258; lean_object* x_259; +x_258 = lean_ctor_get(x_255, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_255, 1); +lean_inc(x_259); +lean_dec(x_255); +x_177 = x_258; +x_178 = x_259; +goto block_195; +} } } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_260; lean_object* x_261; +lean_dec(x_213); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - x_106 = lean_box(0); +x_260 = lean_ctor_get(x_249, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_249, 1); +lean_inc(x_261); +lean_dec(x_249); +x_177 = x_260; +x_178 = x_261; +goto block_195; } -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; } } +else +{ +lean_object* x_262; lean_object* x_263; +lean_dec(x_213); +lean_dec(x_212); +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_262 = lean_ctor_get(x_214, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_214, 1); +lean_inc(x_263); +lean_dec(x_214); +x_177 = x_262; +x_178 = x_263; +goto block_195; +} +block_176: +{ +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; uint8_t x_172; +x_160 = lean_st_ref_get(x_9, x_159); +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +lean_dec(x_160); +x_162 = lean_st_ref_get(x_7, x_161); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = lean_ctor_get(x_156, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_156, 1); +lean_inc(x_166); +lean_dec(x_156); +x_167 = lean_ctor_get(x_163, 2); +lean_inc(x_167); +lean_dec(x_163); +x_168 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +lean_ctor_set(x_168, 2, x_167); +x_169 = lean_st_ref_get(x_9, x_164); +lean_dec(x_9); +x_170 = lean_ctor_get(x_169, 1); +lean_inc(x_170); +lean_dec(x_169); +x_171 = lean_st_ref_set(x_7, x_168, x_170); +lean_dec(x_7); +x_172 = !lean_is_exclusive(x_171); +if (x_172 == 0) +{ +lean_object* x_173; +x_173 = lean_ctor_get(x_171, 0); +lean_dec(x_173); +lean_ctor_set(x_171, 0, x_158); +return x_171; +} +else +{ +lean_object* x_174; lean_object* x_175; +x_174 = lean_ctor_get(x_171, 1); +lean_inc(x_174); +lean_dec(x_171); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_158); +lean_ctor_set(x_175, 1, x_174); +return x_175; +} +} +block_195: +{ +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; uint8_t x_191; +x_179 = lean_st_ref_get(x_9, x_178); +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +lean_dec(x_179); +x_181 = lean_st_ref_get(x_7, x_180); +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 = lean_ctor_get(x_156, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_156, 1); +lean_inc(x_185); +lean_dec(x_156); +x_186 = lean_ctor_get(x_182, 2); +lean_inc(x_186); +lean_dec(x_182); +x_187 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_187, 0, x_184); +lean_ctor_set(x_187, 1, x_185); +lean_ctor_set(x_187, 2, x_186); +x_188 = lean_st_ref_get(x_9, x_183); +lean_dec(x_9); +x_189 = lean_ctor_get(x_188, 1); +lean_inc(x_189); +lean_dec(x_188); +x_190 = lean_st_ref_set(x_7, x_187, x_189); +lean_dec(x_7); +x_191 = !lean_is_exclusive(x_190); +if (x_191 == 0) +{ +lean_object* x_192; +x_192 = lean_ctor_get(x_190, 0); +lean_dec(x_192); +lean_ctor_set_tag(x_190, 1); +lean_ctor_set(x_190, 0, x_177); +return x_190; +} +else +{ +lean_object* x_193; lean_object* x_194; +x_193 = lean_ctor_get(x_190, 1); +lean_inc(x_193); +lean_dec(x_190); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_177); +lean_ctor_set(x_194, 1, x_193); +return x_194; +} } } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_264; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -12906,311 +14061,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_264 = !lean_is_exclusive(x_147); +if (x_264 == 0) { -return x_14; +return x_147; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = lean_ctor_get(x_147, 0); +x_266 = lean_ctor_get(x_147, 1); +lean_inc(x_266); +lean_inc(x_265); +lean_dec(x_147); +x_267 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_267, 0, x_265); +lean_ctor_set(x_267, 1, x_266); +return x_267; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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: +case 3: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); +lean_object* x_268; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +x_268 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_268) == 0) +{ +lean_object* x_269; +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +if (lean_obj_tag(x_269) == 0) +{ +lean_object* x_270; lean_object* x_271; +lean_dec(x_3); +lean_dec(x_1); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +lean_dec(x_268); +x_271 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_270); +return x_271; +} +else +{ +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_298; lean_object* x_299; 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; +x_272 = lean_ctor_get(x_268, 1); +lean_inc(x_272); +lean_dec(x_268); +x_273 = lean_ctor_get(x_269, 0); +lean_inc(x_273); +lean_dec(x_269); +x_274 = lean_st_ref_get(x_9, x_272); +x_275 = lean_ctor_get(x_274, 1); +lean_inc(x_275); +lean_dec(x_274); +x_276 = lean_st_ref_get(x_7, x_275); +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_ctor_get(x_276, 1); +lean_inc(x_278); +lean_dec(x_276); +x_317 = lean_st_ref_get(x_9, x_278); +x_318 = lean_ctor_get(x_317, 1); +lean_inc(x_318); +lean_dec(x_317); +x_319 = lean_st_ref_take(x_7, x_318); +x_320 = lean_ctor_get(x_319, 0); +lean_inc(x_320); +x_321 = lean_ctor_get(x_319, 1); +lean_inc(x_321); +lean_dec(x_319); +x_322 = lean_ctor_get(x_320, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_320, 2); +lean_inc(x_323); +lean_dec(x_320); +x_324 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_325 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_325, 0, x_322); +lean_ctor_set(x_325, 1, x_324); +lean_ctor_set(x_325, 2, x_323); +x_326 = lean_st_ref_set(x_7, x_325, x_321); +x_327 = lean_ctor_get(x_326, 1); +lean_inc(x_327); +lean_dec(x_326); +x_328 = l_Lean_Compiler_visitMatch(x_11, x_273, x_7, x_8, x_9, x_327); +x_329 = lean_ctor_get(x_328, 0); +lean_inc(x_329); +x_330 = lean_ctor_get(x_329, 1); +lean_inc(x_330); +x_331 = lean_ctor_get(x_328, 1); +lean_inc(x_331); +lean_dec(x_328); +x_332 = lean_ctor_get(x_329, 0); +lean_inc(x_332); +lean_dec(x_329); +x_333 = lean_ctor_get(x_330, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_330, 1); +lean_inc(x_334); +lean_dec(x_330); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_335 = lean_apply_7(x_2, x_332, x_5, x_6, x_7, x_8, x_9, x_331); +if (lean_obj_tag(x_335) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_336; lean_object* x_337; lean_object* x_338; uint8_t x_339; +x_336 = lean_ctor_get(x_335, 1); +lean_inc(x_336); +lean_dec(x_335); +x_337 = lean_array_get_size(x_333); +x_338 = lean_unsigned_to_nat(0u); +x_339 = lean_nat_dec_lt(x_338, x_337); +if (x_339 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_340; uint8_t x_341; +lean_dec(x_337); +lean_dec(x_333); +x_340 = lean_array_get_size(x_334); +x_341 = lean_nat_dec_lt(x_338, x_340); +if (x_341 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_342; +lean_dec(x_340); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_342 = lean_box(0); +x_279 = x_342; +x_280 = x_336; +goto block_297; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_343; +x_343 = lean_nat_dec_le(x_340, x_340); +if (x_343 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_344; +lean_dec(x_340); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_344 = lean_box(0); +x_279 = x_344; +x_280 = x_336; +goto block_297; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_345; size_t x_346; lean_object* x_347; lean_object* x_348; +x_345 = 0; +x_346 = lean_usize_of_nat(x_340); +lean_dec(x_340); +x_347 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_348 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(x_1, x_2, x_3, x_334, x_345, x_346, x_347, x_5, x_6, x_7, x_8, x_9, x_336); +lean_dec(x_334); +if (lean_obj_tag(x_348) == 0) +{ +lean_object* x_349; lean_object* x_350; +x_349 = lean_ctor_get(x_348, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_348, 1); +lean_inc(x_350); +lean_dec(x_348); +x_279 = x_349; +x_280 = x_350; +goto block_297; +} +else +{ +lean_object* x_351; lean_object* x_352; +x_351 = lean_ctor_get(x_348, 0); +lean_inc(x_351); +x_352 = lean_ctor_get(x_348, 1); +lean_inc(x_352); +lean_dec(x_348); +x_298 = x_351; +x_299 = x_352; +goto block_316; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_353; +x_353 = lean_nat_dec_le(x_337, x_337); +if (x_353 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_354; uint8_t x_355; +lean_dec(x_337); +lean_dec(x_333); +x_354 = lean_array_get_size(x_334); +x_355 = lean_nat_dec_lt(x_338, x_354); +if (x_355 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_356; +lean_dec(x_354); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_356 = lean_box(0); +x_279 = x_356; +x_280 = x_336; +goto block_297; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_357; +x_357 = lean_nat_dec_le(x_354, x_354); +if (x_357 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_358; +lean_dec(x_354); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_358 = lean_box(0); +x_279 = x_358; +x_280 = x_336; +goto block_297; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_359; size_t x_360; lean_object* x_361; lean_object* x_362; +x_359 = 0; +x_360 = lean_usize_of_nat(x_354); +lean_dec(x_354); +x_361 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_362 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(x_1, x_2, x_3, x_334, x_359, x_360, x_361, x_5, x_6, x_7, x_8, x_9, x_336); +lean_dec(x_334); +if (lean_obj_tag(x_362) == 0) +{ +lean_object* x_363; lean_object* x_364; +x_363 = lean_ctor_get(x_362, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_362, 1); +lean_inc(x_364); +lean_dec(x_362); +x_279 = x_363; +x_280 = x_364; +goto block_297; +} +else +{ +lean_object* x_365; lean_object* x_366; +x_365 = lean_ctor_get(x_362, 0); +lean_inc(x_365); +x_366 = lean_ctor_get(x_362, 1); +lean_inc(x_366); +lean_dec(x_362); +x_298 = x_365; +x_299 = x_366; +goto block_316; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_367; size_t x_368; lean_object* x_369; lean_object* x_370; +x_367 = 0; +x_368 = lean_usize_of_nat(x_337); +lean_dec(x_337); +x_369 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_370 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(x_2, x_333, x_367, x_368, x_369, x_5, x_6, x_7, x_8, x_9, x_336); +lean_dec(x_333); +if (lean_obj_tag(x_370) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_371; lean_object* x_372; uint8_t x_373; +x_371 = lean_ctor_get(x_370, 1); +lean_inc(x_371); +lean_dec(x_370); +x_372 = lean_array_get_size(x_334); +x_373 = lean_nat_dec_lt(x_338, x_372); +if (x_373 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_372); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_279 = x_369; +x_280 = x_371; +goto block_297; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_374; +x_374 = lean_nat_dec_le(x_372, x_372); +if (x_374 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_372); +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_279 = x_369; +x_280 = x_371; +goto block_297; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_375; lean_object* x_376; +x_375 = lean_usize_of_nat(x_372); +lean_dec(x_372); +lean_inc(x_9); +lean_inc(x_7); +x_376 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_1, x_2, x_3, x_334, x_367, x_375, x_369, x_5, x_6, x_7, x_8, x_9, x_371); +lean_dec(x_334); +if (lean_obj_tag(x_376) == 0) +{ +lean_object* x_377; lean_object* x_378; +x_377 = lean_ctor_get(x_376, 0); +lean_inc(x_377); +x_378 = lean_ctor_get(x_376, 1); +lean_inc(x_378); +lean_dec(x_376); +x_279 = x_377; +x_280 = x_378; +goto block_297; +} +else +{ +lean_object* x_379; lean_object* x_380; +x_379 = lean_ctor_get(x_376, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_376, 1); +lean_inc(x_380); +lean_dec(x_376); +x_298 = x_379; +x_299 = x_380; +goto block_316; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_381; lean_object* x_382; +lean_dec(x_334); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_381 = lean_ctor_get(x_370, 0); +lean_inc(x_381); +x_382 = lean_ctor_get(x_370, 1); +lean_inc(x_382); +lean_dec(x_370); +x_298 = x_381; +x_299 = x_382; +goto block_316; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_383; lean_object* x_384; +lean_dec(x_334); +lean_dec(x_333); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_383 = lean_ctor_get(x_335, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_335, 1); +lean_inc(x_384); +lean_dec(x_335); +x_298 = x_383; +x_299 = x_384; +goto block_316; +} +block_297: +{ +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; uint8_t x_293; +x_281 = lean_st_ref_get(x_9, x_280); +x_282 = lean_ctor_get(x_281, 1); +lean_inc(x_282); +lean_dec(x_281); +x_283 = lean_st_ref_get(x_7, x_282); +x_284 = lean_ctor_get(x_283, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_283, 1); +lean_inc(x_285); +lean_dec(x_283); +x_286 = lean_ctor_get(x_277, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_277, 1); +lean_inc(x_287); +lean_dec(x_277); +x_288 = lean_ctor_get(x_284, 2); +lean_inc(x_288); +lean_dec(x_284); +x_289 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_289, 0, x_286); +lean_ctor_set(x_289, 1, x_287); +lean_ctor_set(x_289, 2, x_288); +x_290 = lean_st_ref_get(x_9, x_285); +lean_dec(x_9); +x_291 = lean_ctor_get(x_290, 1); +lean_inc(x_291); +lean_dec(x_290); +x_292 = lean_st_ref_set(x_7, x_289, x_291); +lean_dec(x_7); +x_293 = !lean_is_exclusive(x_292); +if (x_293 == 0) +{ +lean_object* x_294; +x_294 = lean_ctor_get(x_292, 0); +lean_dec(x_294); +lean_ctor_set(x_292, 0, x_279); +return x_292; +} +else +{ +lean_object* x_295; lean_object* x_296; +x_295 = lean_ctor_get(x_292, 1); +lean_inc(x_295); +lean_dec(x_292); +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_279); +lean_ctor_set(x_296, 1, x_295); +return x_296; +} +} +block_316: +{ +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; uint8_t x_312; +x_300 = lean_st_ref_get(x_9, x_299); +x_301 = lean_ctor_get(x_300, 1); +lean_inc(x_301); +lean_dec(x_300); +x_302 = lean_st_ref_get(x_7, x_301); +x_303 = lean_ctor_get(x_302, 0); +lean_inc(x_303); +x_304 = lean_ctor_get(x_302, 1); +lean_inc(x_304); +lean_dec(x_302); +x_305 = lean_ctor_get(x_277, 0); +lean_inc(x_305); +x_306 = lean_ctor_get(x_277, 1); +lean_inc(x_306); +lean_dec(x_277); +x_307 = lean_ctor_get(x_303, 2); +lean_inc(x_307); +lean_dec(x_303); +x_308 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_308, 0, x_305); +lean_ctor_set(x_308, 1, x_306); +lean_ctor_set(x_308, 2, x_307); +x_309 = lean_st_ref_get(x_9, x_304); +lean_dec(x_9); +x_310 = lean_ctor_get(x_309, 1); +lean_inc(x_310); +lean_dec(x_309); +x_311 = lean_st_ref_set(x_7, x_308, x_310); +lean_dec(x_7); +x_312 = !lean_is_exclusive(x_311); +if (x_312 == 0) +{ +lean_object* x_313; +x_313 = lean_ctor_get(x_311, 0); +lean_dec(x_313); +lean_ctor_set_tag(x_311, 1); +lean_ctor_set(x_311, 0, x_298); +return x_311; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_311, 1); +lean_inc(x_314); +lean_dec(x_311); +x_315 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_315, 0, x_298); +lean_ctor_set(x_315, 1, x_314); +return x_315; } } } } else { -uint8_t x_60; -lean_dec(x_13); +uint8_t x_385; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -13218,601 +14600,1077 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_385 = !lean_is_exclusive(x_268); +if (x_385 == 0) { -return x_43; +return x_268; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} +lean_object* x_386; lean_object* x_387; lean_object* x_388; +x_386 = lean_ctor_get(x_268, 0); +x_387 = lean_ctor_get(x_268, 1); +lean_inc(x_387); +lean_inc(x_386); +lean_dec(x_268); +x_388 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_388, 0, x_386); +lean_ctor_set(x_388, 1, x_387); +return x_388; } } } -} -else +case 4: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +lean_object* x_389; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_389 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_389) == 0) +{ +lean_object* x_390; +x_390 = lean_ctor_get(x_389, 0); +lean_inc(x_390); +if (lean_obj_tag(x_390) == 0) +{ +lean_object* x_391; lean_object* x_392; +lean_dec(x_3); +lean_dec(x_1); +x_391 = lean_ctor_get(x_389, 1); +lean_inc(x_391); +lean_dec(x_389); +x_392 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_391); +return x_392; +} +else +{ +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_419; lean_object* x_420; 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_393 = lean_ctor_get(x_389, 1); +lean_inc(x_393); +lean_dec(x_389); +x_394 = lean_ctor_get(x_390, 0); +lean_inc(x_394); +lean_dec(x_390); +x_395 = lean_st_ref_get(x_9, x_393); +x_396 = lean_ctor_get(x_395, 1); +lean_inc(x_396); +lean_dec(x_395); +x_397 = lean_st_ref_get(x_7, x_396); +x_398 = lean_ctor_get(x_397, 0); +lean_inc(x_398); +x_399 = lean_ctor_get(x_397, 1); +lean_inc(x_399); +lean_dec(x_397); +x_438 = lean_st_ref_get(x_9, x_399); +x_439 = lean_ctor_get(x_438, 1); +lean_inc(x_439); +lean_dec(x_438); +x_440 = lean_st_ref_take(x_7, x_439); +x_441 = lean_ctor_get(x_440, 0); +lean_inc(x_441); +x_442 = lean_ctor_get(x_440, 1); +lean_inc(x_442); +lean_dec(x_440); +x_443 = lean_ctor_get(x_441, 0); +lean_inc(x_443); +x_444 = lean_ctor_get(x_441, 2); +lean_inc(x_444); +lean_dec(x_441); +x_445 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_446 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_446, 0, x_443); +lean_ctor_set(x_446, 1, x_445); +lean_ctor_set(x_446, 2, x_444); +x_447 = lean_st_ref_set(x_7, x_446, x_442); +x_448 = lean_ctor_get(x_447, 1); +lean_inc(x_448); +lean_dec(x_447); +x_449 = l_Lean_Compiler_visitMatch(x_11, x_394, x_7, x_8, x_9, x_448); +x_450 = lean_ctor_get(x_449, 0); +lean_inc(x_450); +x_451 = lean_ctor_get(x_450, 1); +lean_inc(x_451); +x_452 = lean_ctor_get(x_449, 1); +lean_inc(x_452); +lean_dec(x_449); +x_453 = lean_ctor_get(x_450, 0); +lean_inc(x_453); +lean_dec(x_450); +x_454 = lean_ctor_get(x_451, 0); +lean_inc(x_454); +x_455 = lean_ctor_get(x_451, 1); +lean_inc(x_455); +lean_dec(x_451); +lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_456 = lean_apply_7(x_2, x_453, x_5, x_6, x_7, x_8, x_9, x_452); +if (lean_obj_tag(x_456) == 0) { -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +lean_object* x_457; lean_object* x_458; lean_object* x_459; uint8_t x_460; +x_457 = lean_ctor_get(x_456, 1); +lean_inc(x_457); +lean_dec(x_456); +x_458 = lean_array_get_size(x_454); +x_459 = lean_unsigned_to_nat(0u); +x_460 = lean_nat_dec_lt(x_459, x_458); +if (x_460 == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_461; uint8_t x_462; +lean_dec(x_458); +lean_dec(x_454); +x_461 = lean_array_get_size(x_455); +x_462 = lean_nat_dec_lt(x_459, x_461); +if (x_462 == 0) +{ +lean_object* x_463; +lean_dec(x_461); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_463 = lean_box(0); +x_400 = x_463; +x_401 = x_457; +goto block_418; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_464; +x_464 = lean_nat_dec_le(x_461, x_461); +if (x_464 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_465; +lean_dec(x_461); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_465 = lean_box(0); +x_400 = x_465; +x_401 = x_457; +goto block_418; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +size_t x_466; size_t x_467; lean_object* x_468; lean_object* x_469; +x_466 = 0; +x_467 = lean_usize_of_nat(x_461); +lean_dec(x_461); +x_468 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_469 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(x_1, x_2, x_3, x_455, x_466, x_467, x_468, x_5, x_6, x_7, x_8, x_9, x_457); +lean_dec(x_455); +if (lean_obj_tag(x_469) == 0) +{ +lean_object* x_470; lean_object* x_471; +x_470 = lean_ctor_get(x_469, 0); +lean_inc(x_470); +x_471 = lean_ctor_get(x_469, 1); +lean_inc(x_471); +lean_dec(x_469); +x_400 = x_470; +x_401 = x_471; +goto block_418; +} +else +{ +lean_object* x_472; lean_object* x_473; +x_472 = lean_ctor_get(x_469, 0); +lean_inc(x_472); +x_473 = lean_ctor_get(x_469, 1); +lean_inc(x_473); +lean_dec(x_469); +x_419 = x_472; +x_420 = x_473; +goto block_437; +} } } } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +uint8_t x_474; +x_474 = lean_nat_dec_le(x_458, x_458); +if (x_474 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +lean_object* x_475; uint8_t x_476; +lean_dec(x_458); +lean_dec(x_454); +x_475 = lean_array_get_size(x_455); +x_476 = lean_nat_dec_lt(x_459, x_475); +if (x_476 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_477; +lean_dec(x_475); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_477 = lean_box(0); +x_400 = x_477; +x_401 = x_457; +goto block_418; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_478; +x_478 = lean_nat_dec_le(x_475, x_475); +if (x_478 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_479; +lean_dec(x_475); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_479 = lean_box(0); +x_400 = x_479; +x_401 = x_457; +goto block_418; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +size_t x_480; size_t x_481; lean_object* x_482; lean_object* x_483; +x_480 = 0; +x_481 = lean_usize_of_nat(x_475); +lean_dec(x_475); +x_482 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_483 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(x_1, x_2, x_3, x_455, x_480, x_481, x_482, x_5, x_6, x_7, x_8, x_9, x_457); +lean_dec(x_455); +if (lean_obj_tag(x_483) == 0) +{ +lean_object* x_484; lean_object* x_485; +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_400 = x_484; +x_401 = x_485; +goto block_418; +} +else +{ +lean_object* x_486; lean_object* x_487; +x_486 = lean_ctor_get(x_483, 0); +lean_inc(x_486); +x_487 = lean_ctor_get(x_483, 1); +lean_inc(x_487); +lean_dec(x_483); +x_419 = x_486; +x_420 = x_487; +goto block_437; +} } } } else { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); +size_t x_488; size_t x_489; lean_object* x_490; lean_object* x_491; +x_488 = 0; +x_489 = lean_usize_of_nat(x_458); +lean_dec(x_458); +x_490 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +lean_inc(x_2); +x_491 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(x_2, x_454, x_488, x_489, x_490, x_5, x_6, x_7, x_8, x_9, x_457); +lean_dec(x_454); +if (lean_obj_tag(x_491) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_492; lean_object* x_493; uint8_t x_494; +x_492 = lean_ctor_get(x_491, 1); +lean_inc(x_492); +lean_dec(x_491); +x_493 = lean_array_get_size(x_455); +x_494 = lean_nat_dec_lt(x_459, x_493); +if (x_494 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_493); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_400 = x_490; +x_401 = x_492; +goto block_418; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_495; +x_495 = lean_nat_dec_le(x_493, x_493); +if (x_495 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_493); +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +x_400 = x_490; +x_401 = x_492; +goto block_418; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +else +{ +size_t x_496; lean_object* x_497; +x_496 = lean_usize_of_nat(x_493); +lean_dec(x_493); +lean_inc(x_9); +lean_inc(x_7); +x_497 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_1, x_2, x_3, x_455, x_488, x_496, x_490, x_5, x_6, x_7, x_8, x_9, x_492); +lean_dec(x_455); +if (lean_obj_tag(x_497) == 0) +{ +lean_object* x_498; lean_object* x_499; +x_498 = lean_ctor_get(x_497, 0); +lean_inc(x_498); +x_499 = lean_ctor_get(x_497, 1); +lean_inc(x_499); +lean_dec(x_497); +x_400 = x_498; +x_401 = x_499; +goto block_418; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_500; lean_object* x_501; +x_500 = lean_ctor_get(x_497, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_497, 1); +lean_inc(x_501); +lean_dec(x_497); +x_419 = x_500; +x_420 = x_501; +goto block_437; +} } } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_502; lean_object* x_503; +lean_dec(x_455); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} +x_502 = lean_ctor_get(x_491, 0); +lean_inc(x_502); +x_503 = lean_ctor_get(x_491, 1); +lean_inc(x_503); +lean_dec(x_491); +x_419 = x_502; +x_420 = x_503; +goto block_437; } } } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_504; lean_object* x_505; +lean_dec(x_455); +lean_dec(x_454); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_504 = lean_ctor_get(x_456, 0); +lean_inc(x_504); +x_505 = lean_ctor_get(x_456, 1); +lean_inc(x_505); +lean_dec(x_456); +x_419 = x_504; +x_420 = x_505; +goto block_437; +} +block_418: { -return x_14; +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; uint8_t x_414; +x_402 = lean_st_ref_get(x_9, x_401); +x_403 = lean_ctor_get(x_402, 1); +lean_inc(x_403); +lean_dec(x_402); +x_404 = lean_st_ref_get(x_7, x_403); +x_405 = lean_ctor_get(x_404, 0); +lean_inc(x_405); +x_406 = lean_ctor_get(x_404, 1); +lean_inc(x_406); +lean_dec(x_404); +x_407 = lean_ctor_get(x_398, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_398, 1); +lean_inc(x_408); +lean_dec(x_398); +x_409 = lean_ctor_get(x_405, 2); +lean_inc(x_409); +lean_dec(x_405); +x_410 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_410, 0, x_407); +lean_ctor_set(x_410, 1, x_408); +lean_ctor_set(x_410, 2, x_409); +x_411 = lean_st_ref_get(x_9, x_406); +lean_dec(x_9); +x_412 = lean_ctor_get(x_411, 1); +lean_inc(x_412); +lean_dec(x_411); +x_413 = lean_st_ref_set(x_7, x_410, x_412); +lean_dec(x_7); +x_414 = !lean_is_exclusive(x_413); +if (x_414 == 0) +{ +lean_object* x_415; +x_415 = lean_ctor_get(x_413, 0); +lean_dec(x_415); +lean_ctor_set(x_413, 0, x_400); +return x_413; +} +else +{ +lean_object* x_416; lean_object* x_417; +x_416 = lean_ctor_get(x_413, 1); +lean_inc(x_416); +lean_dec(x_413); +x_417 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_417, 0, x_400); +lean_ctor_set(x_417, 1, x_416); +return x_417; +} +} +block_437: +{ +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; uint8_t x_433; +x_421 = lean_st_ref_get(x_9, x_420); +x_422 = lean_ctor_get(x_421, 1); +lean_inc(x_422); +lean_dec(x_421); +x_423 = lean_st_ref_get(x_7, x_422); +x_424 = lean_ctor_get(x_423, 0); +lean_inc(x_424); +x_425 = lean_ctor_get(x_423, 1); +lean_inc(x_425); +lean_dec(x_423); +x_426 = lean_ctor_get(x_398, 0); +lean_inc(x_426); +x_427 = lean_ctor_get(x_398, 1); +lean_inc(x_427); +lean_dec(x_398); +x_428 = lean_ctor_get(x_424, 2); +lean_inc(x_428); +lean_dec(x_424); +x_429 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_429, 0, x_426); +lean_ctor_set(x_429, 1, x_427); +lean_ctor_set(x_429, 2, x_428); +x_430 = lean_st_ref_get(x_9, x_425); +lean_dec(x_9); +x_431 = lean_ctor_get(x_430, 1); +lean_inc(x_431); +lean_dec(x_430); +x_432 = lean_st_ref_set(x_7, x_429, x_431); +lean_dec(x_7); +x_433 = !lean_is_exclusive(x_432); +if (x_433 == 0) +{ +lean_object* x_434; +x_434 = lean_ctor_get(x_432, 0); +lean_dec(x_434); +lean_ctor_set_tag(x_432, 1); +lean_ctor_set(x_432, 0, x_419); +return x_432; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; +lean_object* x_435; lean_object* x_436; +x_435 = lean_ctor_get(x_432, 1); +lean_inc(x_435); +lean_dec(x_432); +x_436 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_436, 0, x_419); +lean_ctor_set(x_436, 1, x_435); +return x_436; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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: +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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +uint8_t x_506; +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_3); +lean_dec(x_2); +lean_dec(x_1); +x_506 = !lean_is_exclusive(x_389); +if (x_506 == 0) +{ +return x_389; +} +else +{ +lean_object* x_507; lean_object* x_508; lean_object* x_509; +x_507 = lean_ctor_get(x_389, 0); +x_508 = lean_ctor_get(x_389, 1); +lean_inc(x_508); +lean_inc(x_507); +lean_dec(x_389); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_507); +lean_ctor_set(x_509, 1, x_508); +return x_509; +} +} +} +case 5: +{ +lean_object* x_510; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_510 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_510) == 0) +{ +lean_object* x_511; +x_511 = lean_ctor_get(x_510, 0); +lean_inc(x_511); +if (lean_obj_tag(x_511) == 0) +{ +lean_object* x_512; lean_object* x_513; +lean_dec(x_3); +lean_dec(x_1); +x_512 = lean_ctor_get(x_510, 1); +lean_inc(x_512); +lean_dec(x_510); +x_513 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_512); +return x_513; +} +else +{ +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_540; lean_object* x_541; 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; +x_514 = lean_ctor_get(x_510, 1); +lean_inc(x_514); +lean_dec(x_510); +x_515 = lean_ctor_get(x_511, 0); +lean_inc(x_515); +lean_dec(x_511); +x_516 = lean_st_ref_get(x_9, x_514); +x_517 = lean_ctor_get(x_516, 1); +lean_inc(x_517); +lean_dec(x_516); +x_518 = lean_st_ref_get(x_7, x_517); +x_519 = lean_ctor_get(x_518, 0); +lean_inc(x_519); +x_520 = lean_ctor_get(x_518, 1); +lean_inc(x_520); +lean_dec(x_518); +x_559 = lean_st_ref_get(x_9, x_520); +x_560 = lean_ctor_get(x_559, 1); +lean_inc(x_560); +lean_dec(x_559); +x_561 = lean_st_ref_take(x_7, x_560); +x_562 = lean_ctor_get(x_561, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_561, 1); +lean_inc(x_563); +lean_dec(x_561); +x_564 = lean_ctor_get(x_562, 0); +lean_inc(x_564); +x_565 = lean_ctor_get(x_562, 2); +lean_inc(x_565); +lean_dec(x_562); +x_566 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_567 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_567, 0, x_564); +lean_ctor_set(x_567, 1, x_566); +lean_ctor_set(x_567, 2, x_565); +x_568 = lean_st_ref_set(x_7, x_567, x_563); +x_569 = lean_ctor_get(x_568, 1); +lean_inc(x_569); +lean_dec(x_568); +x_570 = l_Lean_Compiler_visitMatch(x_11, x_515, x_7, x_8, x_9, x_569); +x_571 = lean_ctor_get(x_570, 0); +lean_inc(x_571); +x_572 = lean_ctor_get(x_571, 1); +lean_inc(x_572); +x_573 = lean_ctor_get(x_570, 1); +lean_inc(x_573); +lean_dec(x_570); +x_574 = lean_ctor_get(x_571, 0); +lean_inc(x_574); +lean_dec(x_571); +x_575 = lean_ctor_get(x_572, 0); +lean_inc(x_575); +x_576 = lean_ctor_get(x_572, 1); +lean_inc(x_576); +lean_dec(x_572); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_577 = lean_apply_7(x_2, x_574, x_5, x_6, x_7, x_8, x_9, x_573); +if (lean_obj_tag(x_577) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_578; lean_object* x_579; lean_object* x_580; uint8_t x_581; +x_578 = lean_ctor_get(x_577, 1); +lean_inc(x_578); +lean_dec(x_577); +x_579 = lean_array_get_size(x_575); +x_580 = lean_unsigned_to_nat(0u); +x_581 = lean_nat_dec_lt(x_580, x_579); +if (x_581 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_582; uint8_t x_583; +lean_dec(x_579); +lean_dec(x_575); +x_582 = lean_array_get_size(x_576); +x_583 = lean_nat_dec_lt(x_580, x_582); +if (x_583 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_584; +lean_dec(x_582); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_584 = lean_box(0); +x_521 = x_584; +x_522 = x_578; +goto block_539; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_585; +x_585 = lean_nat_dec_le(x_582, x_582); +if (x_585 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_586; +lean_dec(x_582); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_586 = lean_box(0); +x_521 = x_586; +x_522 = x_578; +goto block_539; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_587; size_t x_588; lean_object* x_589; lean_object* x_590; +x_587 = 0; +x_588 = lean_usize_of_nat(x_582); +lean_dec(x_582); +x_589 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_590 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(x_1, x_2, x_3, x_576, x_587, x_588, x_589, x_5, x_6, x_7, x_8, x_9, x_578); +lean_dec(x_576); +if (lean_obj_tag(x_590) == 0) +{ +lean_object* x_591; lean_object* x_592; +x_591 = lean_ctor_get(x_590, 0); +lean_inc(x_591); +x_592 = lean_ctor_get(x_590, 1); +lean_inc(x_592); +lean_dec(x_590); +x_521 = x_591; +x_522 = x_592; +goto block_539; +} +else +{ +lean_object* x_593; lean_object* x_594; +x_593 = lean_ctor_get(x_590, 0); +lean_inc(x_593); +x_594 = lean_ctor_get(x_590, 1); +lean_inc(x_594); +lean_dec(x_590); +x_540 = x_593; +x_541 = x_594; +goto block_558; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_595; +x_595 = lean_nat_dec_le(x_579, x_579); +if (x_595 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_596; uint8_t x_597; +lean_dec(x_579); +lean_dec(x_575); +x_596 = lean_array_get_size(x_576); +x_597 = lean_nat_dec_lt(x_580, x_596); +if (x_597 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_598; +lean_dec(x_596); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_598 = lean_box(0); +x_521 = x_598; +x_522 = x_578; +goto block_539; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_599; +x_599 = lean_nat_dec_le(x_596, x_596); +if (x_599 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_600; +lean_dec(x_596); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_600 = lean_box(0); +x_521 = x_600; +x_522 = x_578; +goto block_539; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__23(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_601; size_t x_602; lean_object* x_603; lean_object* x_604; +x_601 = 0; +x_602 = lean_usize_of_nat(x_596); +lean_dec(x_596); +x_603 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_604 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(x_1, x_2, x_3, x_576, x_601, x_602, x_603, x_5, x_6, x_7, x_8, x_9, x_578); +lean_dec(x_576); +if (lean_obj_tag(x_604) == 0) +{ +lean_object* x_605; lean_object* x_606; +x_605 = lean_ctor_get(x_604, 0); +lean_inc(x_605); +x_606 = lean_ctor_get(x_604, 1); +lean_inc(x_606); +lean_dec(x_604); +x_521 = x_605; +x_522 = x_606; +goto block_539; +} +else +{ +lean_object* x_607; lean_object* x_608; +x_607 = lean_ctor_get(x_604, 0); +lean_inc(x_607); +x_608 = lean_ctor_get(x_604, 1); +lean_inc(x_608); +lean_dec(x_604); +x_540 = x_607; +x_541 = x_608; +goto block_558; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_609; size_t x_610; lean_object* x_611; lean_object* x_612; +x_609 = 0; +x_610 = lean_usize_of_nat(x_579); +lean_dec(x_579); +x_611 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_612 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(x_2, x_575, x_609, x_610, x_611, x_5, x_6, x_7, x_8, x_9, x_578); +lean_dec(x_575); +if (lean_obj_tag(x_612) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_613; lean_object* x_614; uint8_t x_615; +x_613 = lean_ctor_get(x_612, 1); +lean_inc(x_613); +lean_dec(x_612); +x_614 = lean_array_get_size(x_576); +x_615 = lean_nat_dec_lt(x_580, x_614); +if (x_615 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_614); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_521 = x_611; +x_522 = x_613; +goto block_539; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_616; +x_616 = lean_nat_dec_le(x_614, x_614); +if (x_616 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_614); +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_521 = x_611; +x_522 = x_613; +goto block_539; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_617; lean_object* x_618; +x_617 = lean_usize_of_nat(x_614); +lean_dec(x_614); +lean_inc(x_9); +lean_inc(x_7); +x_618 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_1, x_2, x_3, x_576, x_609, x_617, x_611, x_5, x_6, x_7, x_8, x_9, x_613); +lean_dec(x_576); +if (lean_obj_tag(x_618) == 0) +{ +lean_object* x_619; lean_object* x_620; +x_619 = lean_ctor_get(x_618, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_618, 1); +lean_inc(x_620); +lean_dec(x_618); +x_521 = x_619; +x_522 = x_620; +goto block_539; +} +else +{ +lean_object* x_621; lean_object* x_622; +x_621 = lean_ctor_get(x_618, 0); +lean_inc(x_621); +x_622 = lean_ctor_get(x_618, 1); +lean_inc(x_622); +lean_dec(x_618); +x_540 = x_621; +x_541 = x_622; +goto block_558; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_623; lean_object* x_624; +lean_dec(x_576); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_623 = lean_ctor_get(x_612, 0); +lean_inc(x_623); +x_624 = lean_ctor_get(x_612, 1); +lean_inc(x_624); +lean_dec(x_612); +x_540 = x_623; +x_541 = x_624; +goto block_558; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_625; lean_object* x_626; +lean_dec(x_576); +lean_dec(x_575); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_625 = lean_ctor_get(x_577, 0); +lean_inc(x_625); +x_626 = lean_ctor_get(x_577, 1); +lean_inc(x_626); +lean_dec(x_577); +x_540 = x_625; +x_541 = x_626; +goto block_558; +} +block_539: +{ +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; uint8_t x_535; +x_523 = lean_st_ref_get(x_9, x_522); +x_524 = lean_ctor_get(x_523, 1); +lean_inc(x_524); +lean_dec(x_523); +x_525 = lean_st_ref_get(x_7, x_524); +x_526 = lean_ctor_get(x_525, 0); +lean_inc(x_526); +x_527 = lean_ctor_get(x_525, 1); +lean_inc(x_527); +lean_dec(x_525); +x_528 = lean_ctor_get(x_519, 0); +lean_inc(x_528); +x_529 = lean_ctor_get(x_519, 1); +lean_inc(x_529); +lean_dec(x_519); +x_530 = lean_ctor_get(x_526, 2); +lean_inc(x_530); +lean_dec(x_526); +x_531 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_531, 0, x_528); +lean_ctor_set(x_531, 1, x_529); +lean_ctor_set(x_531, 2, x_530); +x_532 = lean_st_ref_get(x_9, x_527); +lean_dec(x_9); +x_533 = lean_ctor_get(x_532, 1); +lean_inc(x_533); +lean_dec(x_532); +x_534 = lean_st_ref_set(x_7, x_531, x_533); +lean_dec(x_7); +x_535 = !lean_is_exclusive(x_534); +if (x_535 == 0) +{ +lean_object* x_536; +x_536 = lean_ctor_get(x_534, 0); +lean_dec(x_536); +lean_ctor_set(x_534, 0, x_521); +return x_534; +} +else +{ +lean_object* x_537; lean_object* x_538; +x_537 = lean_ctor_get(x_534, 1); +lean_inc(x_537); +lean_dec(x_534); +x_538 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_538, 0, x_521); +lean_ctor_set(x_538, 1, x_537); +return x_538; +} +} +block_558: +{ +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; uint8_t x_554; +x_542 = lean_st_ref_get(x_9, x_541); +x_543 = lean_ctor_get(x_542, 1); +lean_inc(x_543); +lean_dec(x_542); +x_544 = lean_st_ref_get(x_7, x_543); +x_545 = lean_ctor_get(x_544, 0); +lean_inc(x_545); +x_546 = lean_ctor_get(x_544, 1); +lean_inc(x_546); +lean_dec(x_544); +x_547 = lean_ctor_get(x_519, 0); +lean_inc(x_547); +x_548 = lean_ctor_get(x_519, 1); +lean_inc(x_548); +lean_dec(x_519); +x_549 = lean_ctor_get(x_545, 2); +lean_inc(x_549); +lean_dec(x_545); +x_550 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_550, 0, x_547); +lean_ctor_set(x_550, 1, x_548); +lean_ctor_set(x_550, 2, x_549); +x_551 = lean_st_ref_get(x_9, x_546); +lean_dec(x_9); +x_552 = lean_ctor_get(x_551, 1); +lean_inc(x_552); +lean_dec(x_551); +x_553 = lean_st_ref_set(x_7, x_550, x_552); +lean_dec(x_7); +x_554 = !lean_is_exclusive(x_553); +if (x_554 == 0) +{ +lean_object* x_555; +x_555 = lean_ctor_get(x_553, 0); +lean_dec(x_555); +lean_ctor_set_tag(x_553, 1); +lean_ctor_set(x_553, 0, x_540); +return x_553; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +lean_object* x_556; lean_object* x_557; +x_556 = lean_ctor_get(x_553, 1); +lean_inc(x_556); +lean_dec(x_553); +x_557 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_557, 0, x_540); +lean_ctor_set(x_557, 1, x_556); +return x_557; } } } } else { -uint8_t x_60; -lean_dec(x_13); +uint8_t x_627; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -13820,289 +15678,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_627 = !lean_is_exclusive(x_510); +if (x_627 == 0) { -return x_43; +return x_510; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} +lean_object* x_628; lean_object* x_629; lean_object* x_630; +x_628 = lean_ctor_get(x_510, 0); +x_629 = lean_ctor_get(x_510, 1); +lean_inc(x_629); +lean_inc(x_628); +lean_dec(x_510); +x_630 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_630, 0, x_628); +lean_ctor_set(x_630, 1, x_629); +return x_630; } } } -} -else +case 6: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +lean_object* x_631; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_631 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_631) == 0) +{ +lean_object* x_632; +x_632 = lean_ctor_get(x_631, 0); +lean_inc(x_632); +if (lean_obj_tag(x_632) == 0) +{ +lean_object* x_633; lean_object* x_634; +lean_dec(x_3); +lean_dec(x_1); +x_633 = lean_ctor_get(x_631, 1); +lean_inc(x_633); +lean_dec(x_631); +x_634 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_633); +return x_634; +} +else +{ +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_661; lean_object* x_662; 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_635 = lean_ctor_get(x_631, 1); +lean_inc(x_635); +lean_dec(x_631); +x_636 = lean_ctor_get(x_632, 0); +lean_inc(x_636); +lean_dec(x_632); +x_637 = lean_st_ref_get(x_9, x_635); +x_638 = lean_ctor_get(x_637, 1); +lean_inc(x_638); +lean_dec(x_637); +x_639 = lean_st_ref_get(x_7, x_638); +x_640 = lean_ctor_get(x_639, 0); +lean_inc(x_640); +x_641 = lean_ctor_get(x_639, 1); +lean_inc(x_641); +lean_dec(x_639); +x_680 = lean_st_ref_get(x_9, x_641); +x_681 = lean_ctor_get(x_680, 1); +lean_inc(x_681); +lean_dec(x_680); +x_682 = lean_st_ref_take(x_7, x_681); +x_683 = lean_ctor_get(x_682, 0); +lean_inc(x_683); +x_684 = lean_ctor_get(x_682, 1); +lean_inc(x_684); +lean_dec(x_682); +x_685 = lean_ctor_get(x_683, 0); +lean_inc(x_685); +x_686 = lean_ctor_get(x_683, 2); +lean_inc(x_686); +lean_dec(x_683); +x_687 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_688 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_688, 0, x_685); +lean_ctor_set(x_688, 1, x_687); +lean_ctor_set(x_688, 2, x_686); +x_689 = lean_st_ref_set(x_7, x_688, x_684); +x_690 = lean_ctor_get(x_689, 1); +lean_inc(x_690); +lean_dec(x_689); +x_691 = l_Lean_Compiler_visitMatch(x_11, x_636, x_7, x_8, x_9, x_690); +x_692 = lean_ctor_get(x_691, 0); +lean_inc(x_692); +x_693 = lean_ctor_get(x_692, 1); +lean_inc(x_693); +x_694 = lean_ctor_get(x_691, 1); +lean_inc(x_694); +lean_dec(x_691); +x_695 = lean_ctor_get(x_692, 0); +lean_inc(x_695); +lean_dec(x_692); +x_696 = lean_ctor_get(x_693, 0); +lean_inc(x_696); +x_697 = lean_ctor_get(x_693, 1); +lean_inc(x_697); +lean_dec(x_693); +lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_698 = lean_apply_7(x_2, x_695, x_5, x_6, x_7, x_8, x_9, x_694); +if (lean_obj_tag(x_698) == 0) { -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +lean_object* x_699; lean_object* x_700; lean_object* x_701; uint8_t x_702; +x_699 = lean_ctor_get(x_698, 1); +lean_inc(x_699); +lean_dec(x_698); +x_700 = lean_array_get_size(x_696); +x_701 = lean_unsigned_to_nat(0u); +x_702 = lean_nat_dec_lt(x_701, x_700); +if (x_702 == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_703; uint8_t x_704; +lean_dec(x_700); +lean_dec(x_696); +x_703 = lean_array_get_size(x_697); +x_704 = lean_nat_dec_lt(x_701, x_703); +if (x_704 == 0) +{ +lean_object* x_705; +lean_dec(x_703); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_705 = lean_box(0); +x_642 = x_705; +x_643 = x_699; +goto block_660; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_706; +x_706 = lean_nat_dec_le(x_703, x_703); +if (x_706 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_707; +lean_dec(x_703); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_707 = lean_box(0); +x_642 = x_707; +x_643 = x_699; +goto block_660; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +size_t x_708; size_t x_709; lean_object* x_710; lean_object* x_711; +x_708 = 0; +x_709 = lean_usize_of_nat(x_703); +lean_dec(x_703); +x_710 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_711 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(x_1, x_2, x_3, x_697, x_708, x_709, x_710, x_5, x_6, x_7, x_8, x_9, x_699); +lean_dec(x_697); +if (lean_obj_tag(x_711) == 0) +{ +lean_object* x_712; lean_object* x_713; +x_712 = lean_ctor_get(x_711, 0); +lean_inc(x_712); +x_713 = lean_ctor_get(x_711, 1); +lean_inc(x_713); +lean_dec(x_711); +x_642 = x_712; +x_643 = x_713; +goto block_660; +} +else +{ +lean_object* x_714; lean_object* x_715; +x_714 = lean_ctor_get(x_711, 0); +lean_inc(x_714); +x_715 = lean_ctor_get(x_711, 1); +lean_inc(x_715); +lean_dec(x_711); +x_661 = x_714; +x_662 = x_715; +goto block_679; +} } } } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +uint8_t x_716; +x_716 = lean_nat_dec_le(x_700, x_700); +if (x_716 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +lean_object* x_717; uint8_t x_718; +lean_dec(x_700); +lean_dec(x_696); +x_717 = lean_array_get_size(x_697); +x_718 = lean_nat_dec_lt(x_701, x_717); +if (x_718 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_719; +lean_dec(x_717); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_719 = lean_box(0); +x_642 = x_719; +x_643 = x_699; +goto block_660; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_720; +x_720 = lean_nat_dec_le(x_717, x_717); +if (x_720 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_721; +lean_dec(x_717); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_721 = lean_box(0); +x_642 = x_721; +x_643 = x_699; +goto block_660; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__23(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +size_t x_722; size_t x_723; lean_object* x_724; lean_object* x_725; +x_722 = 0; +x_723 = lean_usize_of_nat(x_717); +lean_dec(x_717); +x_724 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_725 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__23(x_1, x_2, x_3, x_697, x_722, x_723, x_724, x_5, x_6, x_7, x_8, x_9, x_699); +lean_dec(x_697); +if (lean_obj_tag(x_725) == 0) +{ +lean_object* x_726; lean_object* x_727; +x_726 = lean_ctor_get(x_725, 0); +lean_inc(x_726); +x_727 = lean_ctor_get(x_725, 1); +lean_inc(x_727); +lean_dec(x_725); +x_642 = x_726; +x_643 = x_727; +goto block_660; +} +else +{ +lean_object* x_728; lean_object* x_729; +x_728 = lean_ctor_get(x_725, 0); +lean_inc(x_728); +x_729 = lean_ctor_get(x_725, 1); +lean_inc(x_729); +lean_dec(x_725); +x_661 = x_728; +x_662 = x_729; +goto block_679; +} } } } else { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); +size_t x_730; size_t x_731; lean_object* x_732; lean_object* x_733; +x_730 = 0; +x_731 = lean_usize_of_nat(x_700); +lean_dec(x_700); +x_732 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +lean_inc(x_2); +x_733 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(x_2, x_696, x_730, x_731, x_732, x_5, x_6, x_7, x_8, x_9, x_699); +lean_dec(x_696); +if (lean_obj_tag(x_733) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_734; lean_object* x_735; uint8_t x_736; +x_734 = lean_ctor_get(x_733, 1); +lean_inc(x_734); +lean_dec(x_733); +x_735 = lean_array_get_size(x_697); +x_736 = lean_nat_dec_lt(x_701, x_735); +if (x_736 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_735); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_642 = x_732; +x_643 = x_734; +goto block_660; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_737; +x_737 = lean_nat_dec_le(x_735, x_735); +if (x_737 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_735); +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +x_642 = x_732; +x_643 = x_734; +goto block_660; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +else +{ +size_t x_738; lean_object* x_739; +x_738 = lean_usize_of_nat(x_735); +lean_dec(x_735); +lean_inc(x_9); +lean_inc(x_7); +x_739 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_1, x_2, x_3, x_697, x_730, x_738, x_732, x_5, x_6, x_7, x_8, x_9, x_734); +lean_dec(x_697); +if (lean_obj_tag(x_739) == 0) +{ +lean_object* x_740; lean_object* x_741; +x_740 = lean_ctor_get(x_739, 0); +lean_inc(x_740); +x_741 = lean_ctor_get(x_739, 1); +lean_inc(x_741); +lean_dec(x_739); +x_642 = x_740; +x_643 = x_741; +goto block_660; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_742; lean_object* x_743; +x_742 = lean_ctor_get(x_739, 0); +lean_inc(x_742); +x_743 = lean_ctor_get(x_739, 1); +lean_inc(x_743); +lean_dec(x_739); +x_661 = x_742; +x_662 = x_743; +goto block_679; +} } } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_744; lean_object* x_745; +lean_dec(x_697); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - x_106 = lean_box(0); +x_744 = lean_ctor_get(x_733, 0); +lean_inc(x_744); +x_745 = lean_ctor_get(x_733, 1); +lean_inc(x_745); +lean_dec(x_733); +x_661 = x_744; +x_662 = x_745; +goto block_679; } -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; } } +else +{ +lean_object* x_746; lean_object* x_747; +lean_dec(x_697); +lean_dec(x_696); +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_746 = lean_ctor_get(x_698, 0); +lean_inc(x_746); +x_747 = lean_ctor_get(x_698, 1); +lean_inc(x_747); +lean_dec(x_698); +x_661 = x_746; +x_662 = x_747; +goto block_679; +} +block_660: +{ +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; +x_644 = lean_st_ref_get(x_9, x_643); +x_645 = lean_ctor_get(x_644, 1); +lean_inc(x_645); +lean_dec(x_644); +x_646 = lean_st_ref_get(x_7, x_645); +x_647 = lean_ctor_get(x_646, 0); +lean_inc(x_647); +x_648 = lean_ctor_get(x_646, 1); +lean_inc(x_648); +lean_dec(x_646); +x_649 = lean_ctor_get(x_640, 0); +lean_inc(x_649); +x_650 = lean_ctor_get(x_640, 1); +lean_inc(x_650); +lean_dec(x_640); +x_651 = lean_ctor_get(x_647, 2); +lean_inc(x_651); +lean_dec(x_647); +x_652 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_652, 0, x_649); +lean_ctor_set(x_652, 1, x_650); +lean_ctor_set(x_652, 2, x_651); +x_653 = lean_st_ref_get(x_9, x_648); +lean_dec(x_9); +x_654 = lean_ctor_get(x_653, 1); +lean_inc(x_654); +lean_dec(x_653); +x_655 = lean_st_ref_set(x_7, x_652, x_654); +lean_dec(x_7); +x_656 = !lean_is_exclusive(x_655); +if (x_656 == 0) +{ +lean_object* x_657; +x_657 = lean_ctor_get(x_655, 0); +lean_dec(x_657); +lean_ctor_set(x_655, 0, x_642); +return x_655; +} +else +{ +lean_object* x_658; lean_object* x_659; +x_658 = lean_ctor_get(x_655, 1); +lean_inc(x_658); +lean_dec(x_655); +x_659 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_659, 0, x_642); +lean_ctor_set(x_659, 1, x_658); +return x_659; +} +} +block_679: +{ +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; uint8_t x_675; +x_663 = lean_st_ref_get(x_9, x_662); +x_664 = lean_ctor_get(x_663, 1); +lean_inc(x_664); +lean_dec(x_663); +x_665 = lean_st_ref_get(x_7, x_664); +x_666 = lean_ctor_get(x_665, 0); +lean_inc(x_666); +x_667 = lean_ctor_get(x_665, 1); +lean_inc(x_667); +lean_dec(x_665); +x_668 = lean_ctor_get(x_640, 0); +lean_inc(x_668); +x_669 = lean_ctor_get(x_640, 1); +lean_inc(x_669); +lean_dec(x_640); +x_670 = lean_ctor_get(x_666, 2); +lean_inc(x_670); +lean_dec(x_666); +x_671 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_671, 0, x_668); +lean_ctor_set(x_671, 1, x_669); +lean_ctor_set(x_671, 2, x_670); +x_672 = lean_st_ref_get(x_9, x_667); +lean_dec(x_9); +x_673 = lean_ctor_get(x_672, 1); +lean_inc(x_673); +lean_dec(x_672); +x_674 = lean_st_ref_set(x_7, x_671, x_673); +lean_dec(x_7); +x_675 = !lean_is_exclusive(x_674); +if (x_675 == 0) +{ +lean_object* x_676; +x_676 = lean_ctor_get(x_674, 0); +lean_dec(x_676); +lean_ctor_set_tag(x_674, 1); +lean_ctor_set(x_674, 0, x_661); +return x_674; +} +else +{ +lean_object* x_677; lean_object* x_678; +x_677 = lean_ctor_get(x_674, 1); +lean_inc(x_677); +lean_dec(x_674); +x_678 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_678, 0, x_661); +lean_ctor_set(x_678, 1, x_677); +return x_678; +} } } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_748; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14110,311 +16217,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_748 = !lean_is_exclusive(x_631); +if (x_748 == 0) { -return x_14; +return x_631; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +lean_object* x_749; lean_object* x_750; lean_object* x_751; +x_749 = lean_ctor_get(x_631, 0); +x_750 = lean_ctor_get(x_631, 1); +lean_inc(x_750); +lean_inc(x_749); +lean_dec(x_631); +x_751 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_751, 0, x_749); +lean_ctor_set(x_751, 1, x_750); +return x_751; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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, lean_object* x_9) { -_start: +case 7: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); +lean_object* x_752; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +x_752 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_752) == 0) +{ +lean_object* x_753; +x_753 = lean_ctor_get(x_752, 0); +lean_inc(x_753); +if (lean_obj_tag(x_753) == 0) +{ +lean_object* x_754; lean_object* x_755; +lean_dec(x_3); +lean_dec(x_1); +x_754 = lean_ctor_get(x_752, 1); +lean_inc(x_754); +lean_dec(x_752); +x_755 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_754); +return x_755; +} +else +{ +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_782; lean_object* x_783; 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; +x_756 = lean_ctor_get(x_752, 1); +lean_inc(x_756); +lean_dec(x_752); +x_757 = lean_ctor_get(x_753, 0); +lean_inc(x_757); +lean_dec(x_753); +x_758 = lean_st_ref_get(x_9, x_756); +x_759 = lean_ctor_get(x_758, 1); +lean_inc(x_759); +lean_dec(x_758); +x_760 = lean_st_ref_get(x_7, x_759); +x_761 = lean_ctor_get(x_760, 0); +lean_inc(x_761); +x_762 = lean_ctor_get(x_760, 1); +lean_inc(x_762); +lean_dec(x_760); +x_801 = lean_st_ref_get(x_9, x_762); +x_802 = lean_ctor_get(x_801, 1); +lean_inc(x_802); +lean_dec(x_801); +x_803 = lean_st_ref_take(x_7, x_802); +x_804 = lean_ctor_get(x_803, 0); +lean_inc(x_804); +x_805 = lean_ctor_get(x_803, 1); +lean_inc(x_805); +lean_dec(x_803); +x_806 = lean_ctor_get(x_804, 0); +lean_inc(x_806); +x_807 = lean_ctor_get(x_804, 2); +lean_inc(x_807); +lean_dec(x_804); +x_808 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_809 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_809, 0, x_806); +lean_ctor_set(x_809, 1, x_808); +lean_ctor_set(x_809, 2, x_807); +x_810 = lean_st_ref_set(x_7, x_809, x_805); +x_811 = lean_ctor_get(x_810, 1); +lean_inc(x_811); +lean_dec(x_810); +x_812 = l_Lean_Compiler_visitMatch(x_11, x_757, x_7, x_8, x_9, x_811); +x_813 = lean_ctor_get(x_812, 0); +lean_inc(x_813); +x_814 = lean_ctor_get(x_813, 1); +lean_inc(x_814); +x_815 = lean_ctor_get(x_812, 1); +lean_inc(x_815); +lean_dec(x_812); +x_816 = lean_ctor_get(x_813, 0); +lean_inc(x_816); +lean_dec(x_813); +x_817 = lean_ctor_get(x_814, 0); +lean_inc(x_817); +x_818 = lean_ctor_get(x_814, 1); +lean_inc(x_818); +lean_dec(x_814); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_819 = lean_apply_7(x_2, x_816, x_5, x_6, x_7, x_8, x_9, x_815); +if (lean_obj_tag(x_819) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_820; lean_object* x_821; lean_object* x_822; uint8_t x_823; +x_820 = lean_ctor_get(x_819, 1); +lean_inc(x_820); +lean_dec(x_819); +x_821 = lean_array_get_size(x_817); +x_822 = lean_unsigned_to_nat(0u); +x_823 = lean_nat_dec_lt(x_822, x_821); +if (x_823 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_824; uint8_t x_825; +lean_dec(x_821); +lean_dec(x_817); +x_824 = lean_array_get_size(x_818); +x_825 = lean_nat_dec_lt(x_822, x_824); +if (x_825 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_826; +lean_dec(x_824); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_826 = lean_box(0); +x_763 = x_826; +x_764 = x_820; +goto block_781; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_827; +x_827 = lean_nat_dec_le(x_824, x_824); +if (x_827 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_828; +lean_dec(x_824); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_828 = lean_box(0); +x_763 = x_828; +x_764 = x_820; +goto block_781; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_829; size_t x_830; lean_object* x_831; lean_object* x_832; +x_829 = 0; +x_830 = lean_usize_of_nat(x_824); +lean_dec(x_824); +x_831 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_832 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(x_1, x_2, x_3, x_818, x_829, x_830, x_831, x_5, x_6, x_7, x_8, x_9, x_820); +lean_dec(x_818); +if (lean_obj_tag(x_832) == 0) +{ +lean_object* x_833; lean_object* x_834; +x_833 = lean_ctor_get(x_832, 0); +lean_inc(x_833); +x_834 = lean_ctor_get(x_832, 1); +lean_inc(x_834); +lean_dec(x_832); +x_763 = x_833; +x_764 = x_834; +goto block_781; +} +else +{ +lean_object* x_835; lean_object* x_836; +x_835 = lean_ctor_get(x_832, 0); +lean_inc(x_835); +x_836 = lean_ctor_get(x_832, 1); +lean_inc(x_836); +lean_dec(x_832); +x_782 = x_835; +x_783 = x_836; +goto block_800; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_837; +x_837 = lean_nat_dec_le(x_821, x_821); +if (x_837 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_838; uint8_t x_839; +lean_dec(x_821); +lean_dec(x_817); +x_838 = lean_array_get_size(x_818); +x_839 = lean_nat_dec_lt(x_822, x_838); +if (x_839 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_840; +lean_dec(x_838); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_840 = lean_box(0); +x_763 = x_840; +x_764 = x_820; +goto block_781; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_841; +x_841 = lean_nat_dec_le(x_838, x_838); +if (x_841 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_842; +lean_dec(x_838); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_842 = lean_box(0); +x_763 = x_842; +x_764 = x_820; +goto block_781; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_843; size_t x_844; lean_object* x_845; lean_object* x_846; +x_843 = 0; +x_844 = lean_usize_of_nat(x_838); +lean_dec(x_838); +x_845 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_846 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(x_1, x_2, x_3, x_818, x_843, x_844, x_845, x_5, x_6, x_7, x_8, x_9, x_820); +lean_dec(x_818); +if (lean_obj_tag(x_846) == 0) +{ +lean_object* x_847; lean_object* x_848; +x_847 = lean_ctor_get(x_846, 0); +lean_inc(x_847); +x_848 = lean_ctor_get(x_846, 1); +lean_inc(x_848); +lean_dec(x_846); +x_763 = x_847; +x_764 = x_848; +goto block_781; +} +else +{ +lean_object* x_849; lean_object* x_850; +x_849 = lean_ctor_get(x_846, 0); +lean_inc(x_849); +x_850 = lean_ctor_get(x_846, 1); +lean_inc(x_850); +lean_dec(x_846); +x_782 = x_849; +x_783 = x_850; +goto block_800; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_851; size_t x_852; lean_object* x_853; lean_object* x_854; +x_851 = 0; +x_852 = lean_usize_of_nat(x_821); +lean_dec(x_821); +x_853 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_854 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(x_2, x_817, x_851, x_852, x_853, x_5, x_6, x_7, x_8, x_9, x_820); +lean_dec(x_817); +if (lean_obj_tag(x_854) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_855; lean_object* x_856; uint8_t x_857; +x_855 = lean_ctor_get(x_854, 1); +lean_inc(x_855); +lean_dec(x_854); +x_856 = lean_array_get_size(x_818); +x_857 = lean_nat_dec_lt(x_822, x_856); +if (x_857 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_856); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_763 = x_853; +x_764 = x_855; +goto block_781; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_858; +x_858 = lean_nat_dec_le(x_856, x_856); +if (x_858 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_856); +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_763 = x_853; +x_764 = x_855; +goto block_781; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_859; lean_object* x_860; +x_859 = lean_usize_of_nat(x_856); +lean_dec(x_856); +lean_inc(x_9); +lean_inc(x_7); +x_860 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_1, x_2, x_3, x_818, x_851, x_859, x_853, x_5, x_6, x_7, x_8, x_9, x_855); +lean_dec(x_818); +if (lean_obj_tag(x_860) == 0) +{ +lean_object* x_861; lean_object* x_862; +x_861 = lean_ctor_get(x_860, 0); +lean_inc(x_861); +x_862 = lean_ctor_get(x_860, 1); +lean_inc(x_862); +lean_dec(x_860); +x_763 = x_861; +x_764 = x_862; +goto block_781; +} +else +{ +lean_object* x_863; lean_object* x_864; +x_863 = lean_ctor_get(x_860, 0); +lean_inc(x_863); +x_864 = lean_ctor_get(x_860, 1); +lean_inc(x_864); +lean_dec(x_860); +x_782 = x_863; +x_783 = x_864; +goto block_800; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_865; lean_object* x_866; +lean_dec(x_818); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_865 = lean_ctor_get(x_854, 0); +lean_inc(x_865); +x_866 = lean_ctor_get(x_854, 1); +lean_inc(x_866); +lean_dec(x_854); +x_782 = x_865; +x_783 = x_866; +goto block_800; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_867; lean_object* x_868; +lean_dec(x_818); +lean_dec(x_817); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_867 = lean_ctor_get(x_819, 0); +lean_inc(x_867); +x_868 = lean_ctor_get(x_819, 1); +lean_inc(x_868); +lean_dec(x_819); +x_782 = x_867; +x_783 = x_868; +goto block_800; +} +block_781: +{ +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; uint8_t x_777; +x_765 = lean_st_ref_get(x_9, x_764); +x_766 = lean_ctor_get(x_765, 1); +lean_inc(x_766); +lean_dec(x_765); +x_767 = lean_st_ref_get(x_7, x_766); +x_768 = lean_ctor_get(x_767, 0); +lean_inc(x_768); +x_769 = lean_ctor_get(x_767, 1); +lean_inc(x_769); +lean_dec(x_767); +x_770 = lean_ctor_get(x_761, 0); +lean_inc(x_770); +x_771 = lean_ctor_get(x_761, 1); +lean_inc(x_771); +lean_dec(x_761); +x_772 = lean_ctor_get(x_768, 2); +lean_inc(x_772); +lean_dec(x_768); +x_773 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_773, 0, x_770); +lean_ctor_set(x_773, 1, x_771); +lean_ctor_set(x_773, 2, x_772); +x_774 = lean_st_ref_get(x_9, x_769); +lean_dec(x_9); +x_775 = lean_ctor_get(x_774, 1); +lean_inc(x_775); +lean_dec(x_774); +x_776 = lean_st_ref_set(x_7, x_773, x_775); +lean_dec(x_7); +x_777 = !lean_is_exclusive(x_776); +if (x_777 == 0) +{ +lean_object* x_778; +x_778 = lean_ctor_get(x_776, 0); +lean_dec(x_778); +lean_ctor_set(x_776, 0, x_763); +return x_776; +} +else +{ +lean_object* x_779; lean_object* x_780; +x_779 = lean_ctor_get(x_776, 1); +lean_inc(x_779); +lean_dec(x_776); +x_780 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_780, 0, x_763); +lean_ctor_set(x_780, 1, x_779); +return x_780; +} +} +block_800: +{ +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; uint8_t x_796; +x_784 = lean_st_ref_get(x_9, x_783); +x_785 = lean_ctor_get(x_784, 1); +lean_inc(x_785); +lean_dec(x_784); +x_786 = lean_st_ref_get(x_7, x_785); +x_787 = lean_ctor_get(x_786, 0); +lean_inc(x_787); +x_788 = lean_ctor_get(x_786, 1); +lean_inc(x_788); +lean_dec(x_786); +x_789 = lean_ctor_get(x_761, 0); +lean_inc(x_789); +x_790 = lean_ctor_get(x_761, 1); +lean_inc(x_790); +lean_dec(x_761); +x_791 = lean_ctor_get(x_787, 2); +lean_inc(x_791); +lean_dec(x_787); +x_792 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_792, 0, x_789); +lean_ctor_set(x_792, 1, x_790); +lean_ctor_set(x_792, 2, x_791); +x_793 = lean_st_ref_get(x_9, x_788); +lean_dec(x_9); +x_794 = lean_ctor_get(x_793, 1); +lean_inc(x_794); +lean_dec(x_793); +x_795 = lean_st_ref_set(x_7, x_792, x_794); +lean_dec(x_7); +x_796 = !lean_is_exclusive(x_795); +if (x_796 == 0) +{ +lean_object* x_797; +x_797 = lean_ctor_get(x_795, 0); +lean_dec(x_797); +lean_ctor_set_tag(x_795, 1); +lean_ctor_set(x_795, 0, x_782); +return x_795; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +lean_object* x_798; lean_object* x_799; +x_798 = lean_ctor_get(x_795, 1); +lean_inc(x_798); +lean_dec(x_795); +x_799 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_799, 0, x_782); +lean_ctor_set(x_799, 1, x_798); +return x_799; } } } } else { -uint8_t x_60; -lean_dec(x_13); +uint8_t x_869; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14422,289 +16756,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_869 = !lean_is_exclusive(x_752); +if (x_869 == 0) { -return x_43; +return x_752; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} +lean_object* x_870; lean_object* x_871; lean_object* x_872; +x_870 = lean_ctor_get(x_752, 0); +x_871 = lean_ctor_get(x_752, 1); +lean_inc(x_871); +lean_inc(x_870); +lean_dec(x_752); +x_872 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_872, 0, x_870); +lean_ctor_set(x_872, 1, x_871); +return x_872; } } } -} -else +case 8: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +lean_object* x_873; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_873 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_873) == 0) +{ +lean_object* x_874; +x_874 = lean_ctor_get(x_873, 0); +lean_inc(x_874); +if (lean_obj_tag(x_874) == 0) +{ +lean_object* x_875; lean_object* x_876; +lean_dec(x_3); +lean_dec(x_1); +x_875 = lean_ctor_get(x_873, 1); +lean_inc(x_875); +lean_dec(x_873); +x_876 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_875); +return x_876; +} +else +{ +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_903; lean_object* x_904; 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; +x_877 = lean_ctor_get(x_873, 1); +lean_inc(x_877); +lean_dec(x_873); +x_878 = lean_ctor_get(x_874, 0); +lean_inc(x_878); +lean_dec(x_874); +x_879 = lean_st_ref_get(x_9, x_877); +x_880 = lean_ctor_get(x_879, 1); +lean_inc(x_880); +lean_dec(x_879); +x_881 = lean_st_ref_get(x_7, x_880); +x_882 = lean_ctor_get(x_881, 0); +lean_inc(x_882); +x_883 = lean_ctor_get(x_881, 1); +lean_inc(x_883); +lean_dec(x_881); +x_922 = lean_st_ref_get(x_9, x_883); +x_923 = lean_ctor_get(x_922, 1); +lean_inc(x_923); +lean_dec(x_922); +x_924 = lean_st_ref_take(x_7, x_923); +x_925 = lean_ctor_get(x_924, 0); +lean_inc(x_925); +x_926 = lean_ctor_get(x_924, 1); +lean_inc(x_926); +lean_dec(x_924); +x_927 = lean_ctor_get(x_925, 0); +lean_inc(x_927); +x_928 = lean_ctor_get(x_925, 2); +lean_inc(x_928); +lean_dec(x_925); +x_929 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_930 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_930, 0, x_927); +lean_ctor_set(x_930, 1, x_929); +lean_ctor_set(x_930, 2, x_928); +x_931 = lean_st_ref_set(x_7, x_930, x_926); +x_932 = lean_ctor_get(x_931, 1); +lean_inc(x_932); +lean_dec(x_931); +x_933 = l_Lean_Compiler_visitMatch(x_11, x_878, x_7, x_8, x_9, x_932); +x_934 = lean_ctor_get(x_933, 0); +lean_inc(x_934); +x_935 = lean_ctor_get(x_934, 1); +lean_inc(x_935); +x_936 = lean_ctor_get(x_933, 1); +lean_inc(x_936); +lean_dec(x_933); +x_937 = lean_ctor_get(x_934, 0); +lean_inc(x_937); +lean_dec(x_934); +x_938 = lean_ctor_get(x_935, 0); +lean_inc(x_938); +x_939 = lean_ctor_get(x_935, 1); +lean_inc(x_939); +lean_dec(x_935); +lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_940 = lean_apply_7(x_2, x_937, x_5, x_6, x_7, x_8, x_9, x_936); +if (lean_obj_tag(x_940) == 0) { -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +lean_object* x_941; lean_object* x_942; lean_object* x_943; uint8_t x_944; +x_941 = lean_ctor_get(x_940, 1); +lean_inc(x_941); +lean_dec(x_940); +x_942 = lean_array_get_size(x_938); +x_943 = lean_unsigned_to_nat(0u); +x_944 = lean_nat_dec_lt(x_943, x_942); +if (x_944 == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_945; uint8_t x_946; +lean_dec(x_942); +lean_dec(x_938); +x_945 = lean_array_get_size(x_939); +x_946 = lean_nat_dec_lt(x_943, x_945); +if (x_946 == 0) +{ +lean_object* x_947; +lean_dec(x_945); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_947 = lean_box(0); +x_884 = x_947; +x_885 = x_941; +goto block_902; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_948; +x_948 = lean_nat_dec_le(x_945, x_945); +if (x_948 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_949; +lean_dec(x_945); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_949 = lean_box(0); +x_884 = x_949; +x_885 = x_941; +goto block_902; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +size_t x_950; size_t x_951; lean_object* x_952; lean_object* x_953; +x_950 = 0; +x_951 = lean_usize_of_nat(x_945); +lean_dec(x_945); +x_952 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_953 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(x_1, x_2, x_3, x_939, x_950, x_951, x_952, x_5, x_6, x_7, x_8, x_9, x_941); +lean_dec(x_939); +if (lean_obj_tag(x_953) == 0) +{ +lean_object* x_954; lean_object* x_955; +x_954 = lean_ctor_get(x_953, 0); +lean_inc(x_954); +x_955 = lean_ctor_get(x_953, 1); +lean_inc(x_955); +lean_dec(x_953); +x_884 = x_954; +x_885 = x_955; +goto block_902; +} +else +{ +lean_object* x_956; lean_object* x_957; +x_956 = lean_ctor_get(x_953, 0); +lean_inc(x_956); +x_957 = lean_ctor_get(x_953, 1); +lean_inc(x_957); +lean_dec(x_953); +x_903 = x_956; +x_904 = x_957; +goto block_921; +} } } } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +uint8_t x_958; +x_958 = lean_nat_dec_le(x_942, x_942); +if (x_958 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +lean_object* x_959; uint8_t x_960; +lean_dec(x_942); +lean_dec(x_938); +x_959 = lean_array_get_size(x_939); +x_960 = lean_nat_dec_lt(x_943, x_959); +if (x_960 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_961; +lean_dec(x_959); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_961 = lean_box(0); +x_884 = x_961; +x_885 = x_941; +goto block_902; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_962; +x_962 = lean_nat_dec_le(x_959, x_959); +if (x_962 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_963; +lean_dec(x_959); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_963 = lean_box(0); +x_884 = x_963; +x_885 = x_941; +goto block_902; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +size_t x_964; size_t x_965; lean_object* x_966; lean_object* x_967; +x_964 = 0; +x_965 = lean_usize_of_nat(x_959); +lean_dec(x_959); +x_966 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_967 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(x_1, x_2, x_3, x_939, x_964, x_965, x_966, x_5, x_6, x_7, x_8, x_9, x_941); +lean_dec(x_939); +if (lean_obj_tag(x_967) == 0) +{ +lean_object* x_968; lean_object* x_969; +x_968 = lean_ctor_get(x_967, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_967, 1); +lean_inc(x_969); +lean_dec(x_967); +x_884 = x_968; +x_885 = x_969; +goto block_902; +} +else +{ +lean_object* x_970; lean_object* x_971; +x_970 = lean_ctor_get(x_967, 0); +lean_inc(x_970); +x_971 = lean_ctor_get(x_967, 1); +lean_inc(x_971); +lean_dec(x_967); +x_903 = x_970; +x_904 = x_971; +goto block_921; +} } } } else { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); +size_t x_972; size_t x_973; lean_object* x_974; lean_object* x_975; +x_972 = 0; +x_973 = lean_usize_of_nat(x_942); +lean_dec(x_942); +x_974 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +lean_inc(x_2); +x_975 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(x_2, x_938, x_972, x_973, x_974, x_5, x_6, x_7, x_8, x_9, x_941); +lean_dec(x_938); +if (lean_obj_tag(x_975) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_976; lean_object* x_977; uint8_t x_978; +x_976 = lean_ctor_get(x_975, 1); +lean_inc(x_976); +lean_dec(x_975); +x_977 = lean_array_get_size(x_939); +x_978 = lean_nat_dec_lt(x_943, x_977); +if (x_978 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_977); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_884 = x_974; +x_885 = x_976; +goto block_902; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_979; +x_979 = lean_nat_dec_le(x_977, x_977); +if (x_979 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_977); +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +x_884 = x_974; +x_885 = x_976; +goto block_902; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +else +{ +size_t x_980; lean_object* x_981; +x_980 = lean_usize_of_nat(x_977); +lean_dec(x_977); +lean_inc(x_9); +lean_inc(x_7); +x_981 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_1, x_2, x_3, x_939, x_972, x_980, x_974, x_5, x_6, x_7, x_8, x_9, x_976); +lean_dec(x_939); +if (lean_obj_tag(x_981) == 0) +{ +lean_object* x_982; lean_object* x_983; +x_982 = lean_ctor_get(x_981, 0); +lean_inc(x_982); +x_983 = lean_ctor_get(x_981, 1); +lean_inc(x_983); +lean_dec(x_981); +x_884 = x_982; +x_885 = x_983; +goto block_902; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_984; lean_object* x_985; +x_984 = lean_ctor_get(x_981, 0); +lean_inc(x_984); +x_985 = lean_ctor_get(x_981, 1); +lean_inc(x_985); +lean_dec(x_981); +x_903 = x_984; +x_904 = x_985; +goto block_921; +} } } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_986; lean_object* x_987; +lean_dec(x_939); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - x_106 = lean_box(0); +x_986 = lean_ctor_get(x_975, 0); +lean_inc(x_986); +x_987 = lean_ctor_get(x_975, 1); +lean_inc(x_987); +lean_dec(x_975); +x_903 = x_986; +x_904 = x_987; +goto block_921; } -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; } } +else +{ +lean_object* x_988; lean_object* x_989; +lean_dec(x_939); +lean_dec(x_938); +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_988 = lean_ctor_get(x_940, 0); +lean_inc(x_988); +x_989 = lean_ctor_get(x_940, 1); +lean_inc(x_989); +lean_dec(x_940); +x_903 = x_988; +x_904 = x_989; +goto block_921; +} +block_902: +{ +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; uint8_t x_898; +x_886 = lean_st_ref_get(x_9, x_885); +x_887 = lean_ctor_get(x_886, 1); +lean_inc(x_887); +lean_dec(x_886); +x_888 = lean_st_ref_get(x_7, x_887); +x_889 = lean_ctor_get(x_888, 0); +lean_inc(x_889); +x_890 = lean_ctor_get(x_888, 1); +lean_inc(x_890); +lean_dec(x_888); +x_891 = lean_ctor_get(x_882, 0); +lean_inc(x_891); +x_892 = lean_ctor_get(x_882, 1); +lean_inc(x_892); +lean_dec(x_882); +x_893 = lean_ctor_get(x_889, 2); +lean_inc(x_893); +lean_dec(x_889); +x_894 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_894, 0, x_891); +lean_ctor_set(x_894, 1, x_892); +lean_ctor_set(x_894, 2, x_893); +x_895 = lean_st_ref_get(x_9, x_890); +lean_dec(x_9); +x_896 = lean_ctor_get(x_895, 1); +lean_inc(x_896); +lean_dec(x_895); +x_897 = lean_st_ref_set(x_7, x_894, x_896); +lean_dec(x_7); +x_898 = !lean_is_exclusive(x_897); +if (x_898 == 0) +{ +lean_object* x_899; +x_899 = lean_ctor_get(x_897, 0); +lean_dec(x_899); +lean_ctor_set(x_897, 0, x_884); +return x_897; +} +else +{ +lean_object* x_900; lean_object* x_901; +x_900 = lean_ctor_get(x_897, 1); +lean_inc(x_900); +lean_dec(x_897); +x_901 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_901, 0, x_884); +lean_ctor_set(x_901, 1, x_900); +return x_901; +} +} +block_921: +{ +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; uint8_t x_917; +x_905 = lean_st_ref_get(x_9, x_904); +x_906 = lean_ctor_get(x_905, 1); +lean_inc(x_906); +lean_dec(x_905); +x_907 = lean_st_ref_get(x_7, x_906); +x_908 = lean_ctor_get(x_907, 0); +lean_inc(x_908); +x_909 = lean_ctor_get(x_907, 1); +lean_inc(x_909); +lean_dec(x_907); +x_910 = lean_ctor_get(x_882, 0); +lean_inc(x_910); +x_911 = lean_ctor_get(x_882, 1); +lean_inc(x_911); +lean_dec(x_882); +x_912 = lean_ctor_get(x_908, 2); +lean_inc(x_912); +lean_dec(x_908); +x_913 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_913, 0, x_910); +lean_ctor_set(x_913, 1, x_911); +lean_ctor_set(x_913, 2, x_912); +x_914 = lean_st_ref_get(x_9, x_909); +lean_dec(x_9); +x_915 = lean_ctor_get(x_914, 1); +lean_inc(x_915); +lean_dec(x_914); +x_916 = lean_st_ref_set(x_7, x_913, x_915); +lean_dec(x_7); +x_917 = !lean_is_exclusive(x_916); +if (x_917 == 0) +{ +lean_object* x_918; +x_918 = lean_ctor_get(x_916, 0); +lean_dec(x_918); +lean_ctor_set_tag(x_916, 1); +lean_ctor_set(x_916, 0, x_903); +return x_916; +} +else +{ +lean_object* x_919; lean_object* x_920; +x_919 = lean_ctor_get(x_916, 1); +lean_inc(x_919); +lean_dec(x_916); +x_920 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_920, 0, x_903); +lean_ctor_set(x_920, 1, x_919); +return x_920; +} } } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_990; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14712,311 +17295,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_990 = !lean_is_exclusive(x_873); +if (x_990 == 0) { -return x_14; +return x_873; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +lean_object* x_991; lean_object* x_992; lean_object* x_993; +x_991 = lean_ctor_get(x_873, 0); +x_992 = lean_ctor_get(x_873, 1); +lean_inc(x_992); +lean_inc(x_991); +lean_dec(x_873); +x_993 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_993, 0, x_991); +lean_ctor_set(x_993, 1, x_992); +return x_993; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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: +case 9: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); +lean_object* x_994; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +x_994 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_994) == 0) +{ +lean_object* x_995; +x_995 = lean_ctor_get(x_994, 0); +lean_inc(x_995); +if (lean_obj_tag(x_995) == 0) +{ +lean_object* x_996; lean_object* x_997; +lean_dec(x_3); +lean_dec(x_1); +x_996 = lean_ctor_get(x_994, 1); +lean_inc(x_996); +lean_dec(x_994); +x_997 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_996); +return x_997; +} +else +{ +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_1024; lean_object* x_1025; 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; +x_998 = lean_ctor_get(x_994, 1); +lean_inc(x_998); +lean_dec(x_994); +x_999 = lean_ctor_get(x_995, 0); +lean_inc(x_999); +lean_dec(x_995); +x_1000 = lean_st_ref_get(x_9, x_998); +x_1001 = lean_ctor_get(x_1000, 1); +lean_inc(x_1001); +lean_dec(x_1000); +x_1002 = lean_st_ref_get(x_7, x_1001); +x_1003 = lean_ctor_get(x_1002, 0); +lean_inc(x_1003); +x_1004 = lean_ctor_get(x_1002, 1); +lean_inc(x_1004); +lean_dec(x_1002); +x_1043 = lean_st_ref_get(x_9, x_1004); +x_1044 = lean_ctor_get(x_1043, 1); +lean_inc(x_1044); +lean_dec(x_1043); +x_1045 = lean_st_ref_take(x_7, x_1044); +x_1046 = lean_ctor_get(x_1045, 0); +lean_inc(x_1046); +x_1047 = lean_ctor_get(x_1045, 1); +lean_inc(x_1047); +lean_dec(x_1045); +x_1048 = lean_ctor_get(x_1046, 0); +lean_inc(x_1048); +x_1049 = lean_ctor_get(x_1046, 2); +lean_inc(x_1049); +lean_dec(x_1046); +x_1050 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1051 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1051, 0, x_1048); +lean_ctor_set(x_1051, 1, x_1050); +lean_ctor_set(x_1051, 2, x_1049); +x_1052 = lean_st_ref_set(x_7, x_1051, x_1047); +x_1053 = lean_ctor_get(x_1052, 1); +lean_inc(x_1053); +lean_dec(x_1052); +x_1054 = l_Lean_Compiler_visitMatch(x_11, x_999, x_7, x_8, x_9, x_1053); +x_1055 = lean_ctor_get(x_1054, 0); +lean_inc(x_1055); +x_1056 = lean_ctor_get(x_1055, 1); +lean_inc(x_1056); +x_1057 = lean_ctor_get(x_1054, 1); +lean_inc(x_1057); +lean_dec(x_1054); +x_1058 = lean_ctor_get(x_1055, 0); +lean_inc(x_1058); +lean_dec(x_1055); +x_1059 = lean_ctor_get(x_1056, 0); +lean_inc(x_1059); +x_1060 = lean_ctor_get(x_1056, 1); +lean_inc(x_1060); +lean_dec(x_1056); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_1061 = lean_apply_7(x_2, x_1058, x_5, x_6, x_7, x_8, x_9, x_1057); +if (lean_obj_tag(x_1061) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; uint8_t x_1065; +x_1062 = lean_ctor_get(x_1061, 1); +lean_inc(x_1062); +lean_dec(x_1061); +x_1063 = lean_array_get_size(x_1059); +x_1064 = lean_unsigned_to_nat(0u); +x_1065 = lean_nat_dec_lt(x_1064, x_1063); +if (x_1065 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_1066; uint8_t x_1067; +lean_dec(x_1063); +lean_dec(x_1059); +x_1066 = lean_array_get_size(x_1060); +x_1067 = lean_nat_dec_lt(x_1064, x_1066); +if (x_1067 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_1068; +lean_dec(x_1066); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_1068 = lean_box(0); +x_1005 = x_1068; +x_1006 = x_1062; +goto block_1023; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_1069; +x_1069 = lean_nat_dec_le(x_1066, x_1066); +if (x_1069 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_1070; +lean_dec(x_1066); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_1070 = lean_box(0); +x_1005 = x_1070; +x_1006 = x_1062; +goto block_1023; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_1071; size_t x_1072; lean_object* x_1073; lean_object* x_1074; +x_1071 = 0; +x_1072 = lean_usize_of_nat(x_1066); +lean_dec(x_1066); +x_1073 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1074 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(x_1, x_2, x_3, x_1060, x_1071, x_1072, x_1073, x_5, x_6, x_7, x_8, x_9, x_1062); +lean_dec(x_1060); +if (lean_obj_tag(x_1074) == 0) +{ +lean_object* x_1075; lean_object* x_1076; +x_1075 = lean_ctor_get(x_1074, 0); +lean_inc(x_1075); +x_1076 = lean_ctor_get(x_1074, 1); +lean_inc(x_1076); +lean_dec(x_1074); +x_1005 = x_1075; +x_1006 = x_1076; +goto block_1023; +} +else +{ +lean_object* x_1077; lean_object* x_1078; +x_1077 = lean_ctor_get(x_1074, 0); +lean_inc(x_1077); +x_1078 = lean_ctor_get(x_1074, 1); +lean_inc(x_1078); +lean_dec(x_1074); +x_1024 = x_1077; +x_1025 = x_1078; +goto block_1042; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_1079; +x_1079 = lean_nat_dec_le(x_1063, x_1063); +if (x_1079 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_1080; uint8_t x_1081; +lean_dec(x_1063); +lean_dec(x_1059); +x_1080 = lean_array_get_size(x_1060); +x_1081 = lean_nat_dec_lt(x_1064, x_1080); +if (x_1081 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_1082; +lean_dec(x_1080); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_1082 = lean_box(0); +x_1005 = x_1082; +x_1006 = x_1062; +goto block_1023; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_1083; +x_1083 = lean_nat_dec_le(x_1080, x_1080); +if (x_1083 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_1084; +lean_dec(x_1080); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_1084 = lean_box(0); +x_1005 = x_1084; +x_1006 = x_1062; +goto block_1023; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_1085; size_t x_1086; lean_object* x_1087; lean_object* x_1088; +x_1085 = 0; +x_1086 = lean_usize_of_nat(x_1080); +lean_dec(x_1080); +x_1087 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1088 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(x_1, x_2, x_3, x_1060, x_1085, x_1086, x_1087, x_5, x_6, x_7, x_8, x_9, x_1062); +lean_dec(x_1060); +if (lean_obj_tag(x_1088) == 0) +{ +lean_object* x_1089; lean_object* x_1090; +x_1089 = lean_ctor_get(x_1088, 0); +lean_inc(x_1089); +x_1090 = lean_ctor_get(x_1088, 1); +lean_inc(x_1090); +lean_dec(x_1088); +x_1005 = x_1089; +x_1006 = x_1090; +goto block_1023; +} +else +{ +lean_object* x_1091; lean_object* x_1092; +x_1091 = lean_ctor_get(x_1088, 0); +lean_inc(x_1091); +x_1092 = lean_ctor_get(x_1088, 1); +lean_inc(x_1092); +lean_dec(x_1088); +x_1024 = x_1091; +x_1025 = x_1092; +goto block_1042; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_1093; size_t x_1094; lean_object* x_1095; lean_object* x_1096; +x_1093 = 0; +x_1094 = lean_usize_of_nat(x_1063); +lean_dec(x_1063); +x_1095 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_1096 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(x_2, x_1059, x_1093, x_1094, x_1095, x_5, x_6, x_7, x_8, x_9, x_1062); +lean_dec(x_1059); +if (lean_obj_tag(x_1096) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_1097; lean_object* x_1098; uint8_t x_1099; +x_1097 = lean_ctor_get(x_1096, 1); +lean_inc(x_1097); +lean_dec(x_1096); +x_1098 = lean_array_get_size(x_1060); +x_1099 = lean_nat_dec_lt(x_1064, x_1098); +if (x_1099 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_1098); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_1005 = x_1095; +x_1006 = x_1097; +goto block_1023; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_1100; +x_1100 = lean_nat_dec_le(x_1098, x_1098); +if (x_1100 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_1098); +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_1005 = x_1095; +x_1006 = x_1097; +goto block_1023; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_1101; lean_object* x_1102; +x_1101 = lean_usize_of_nat(x_1098); +lean_dec(x_1098); +lean_inc(x_9); +lean_inc(x_7); +x_1102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_1, x_2, x_3, x_1060, x_1093, x_1101, x_1095, x_5, x_6, x_7, x_8, x_9, x_1097); +lean_dec(x_1060); +if (lean_obj_tag(x_1102) == 0) +{ +lean_object* x_1103; lean_object* x_1104; +x_1103 = lean_ctor_get(x_1102, 0); +lean_inc(x_1103); +x_1104 = lean_ctor_get(x_1102, 1); +lean_inc(x_1104); +lean_dec(x_1102); +x_1005 = x_1103; +x_1006 = x_1104; +goto block_1023; +} +else +{ +lean_object* x_1105; lean_object* x_1106; +x_1105 = lean_ctor_get(x_1102, 0); +lean_inc(x_1105); +x_1106 = lean_ctor_get(x_1102, 1); +lean_inc(x_1106); +lean_dec(x_1102); +x_1024 = x_1105; +x_1025 = x_1106; +goto block_1042; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_1107; lean_object* x_1108; +lean_dec(x_1060); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_1107 = lean_ctor_get(x_1096, 0); +lean_inc(x_1107); +x_1108 = lean_ctor_get(x_1096, 1); +lean_inc(x_1108); +lean_dec(x_1096); +x_1024 = x_1107; +x_1025 = x_1108; +goto block_1042; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_1109; lean_object* x_1110; +lean_dec(x_1060); +lean_dec(x_1059); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_1109 = lean_ctor_get(x_1061, 0); +lean_inc(x_1109); +x_1110 = lean_ctor_get(x_1061, 1); +lean_inc(x_1110); +lean_dec(x_1061); +x_1024 = x_1109; +x_1025 = x_1110; +goto block_1042; +} +block_1023: +{ +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; uint8_t x_1019; +x_1007 = lean_st_ref_get(x_9, x_1006); +x_1008 = lean_ctor_get(x_1007, 1); +lean_inc(x_1008); +lean_dec(x_1007); +x_1009 = lean_st_ref_get(x_7, x_1008); +x_1010 = lean_ctor_get(x_1009, 0); +lean_inc(x_1010); +x_1011 = lean_ctor_get(x_1009, 1); +lean_inc(x_1011); +lean_dec(x_1009); +x_1012 = lean_ctor_get(x_1003, 0); +lean_inc(x_1012); +x_1013 = lean_ctor_get(x_1003, 1); +lean_inc(x_1013); +lean_dec(x_1003); +x_1014 = lean_ctor_get(x_1010, 2); +lean_inc(x_1014); +lean_dec(x_1010); +x_1015 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1015, 0, x_1012); +lean_ctor_set(x_1015, 1, x_1013); +lean_ctor_set(x_1015, 2, x_1014); +x_1016 = lean_st_ref_get(x_9, x_1011); +lean_dec(x_9); +x_1017 = lean_ctor_get(x_1016, 1); +lean_inc(x_1017); +lean_dec(x_1016); +x_1018 = lean_st_ref_set(x_7, x_1015, x_1017); +lean_dec(x_7); +x_1019 = !lean_is_exclusive(x_1018); +if (x_1019 == 0) +{ +lean_object* x_1020; +x_1020 = lean_ctor_get(x_1018, 0); +lean_dec(x_1020); +lean_ctor_set(x_1018, 0, x_1005); +return x_1018; +} +else +{ +lean_object* x_1021; lean_object* x_1022; +x_1021 = lean_ctor_get(x_1018, 1); +lean_inc(x_1021); +lean_dec(x_1018); +x_1022 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1022, 0, x_1005); +lean_ctor_set(x_1022, 1, x_1021); +return x_1022; +} +} +block_1042: +{ +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; uint8_t x_1038; +x_1026 = lean_st_ref_get(x_9, x_1025); +x_1027 = lean_ctor_get(x_1026, 1); +lean_inc(x_1027); +lean_dec(x_1026); +x_1028 = lean_st_ref_get(x_7, x_1027); +x_1029 = lean_ctor_get(x_1028, 0); +lean_inc(x_1029); +x_1030 = lean_ctor_get(x_1028, 1); +lean_inc(x_1030); +lean_dec(x_1028); +x_1031 = lean_ctor_get(x_1003, 0); +lean_inc(x_1031); +x_1032 = lean_ctor_get(x_1003, 1); +lean_inc(x_1032); +lean_dec(x_1003); +x_1033 = lean_ctor_get(x_1029, 2); +lean_inc(x_1033); +lean_dec(x_1029); +x_1034 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1034, 0, x_1031); +lean_ctor_set(x_1034, 1, x_1032); +lean_ctor_set(x_1034, 2, x_1033); +x_1035 = lean_st_ref_get(x_9, x_1030); +lean_dec(x_9); +x_1036 = lean_ctor_get(x_1035, 1); +lean_inc(x_1036); +lean_dec(x_1035); +x_1037 = lean_st_ref_set(x_7, x_1034, x_1036); +lean_dec(x_7); +x_1038 = !lean_is_exclusive(x_1037); +if (x_1038 == 0) +{ +lean_object* x_1039; +x_1039 = lean_ctor_get(x_1037, 0); +lean_dec(x_1039); +lean_ctor_set_tag(x_1037, 1); +lean_ctor_set(x_1037, 0, x_1024); +return x_1037; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +lean_object* x_1040; lean_object* x_1041; +x_1040 = lean_ctor_get(x_1037, 1); +lean_inc(x_1040); +lean_dec(x_1037); +x_1041 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1041, 0, x_1024); +lean_ctor_set(x_1041, 1, x_1040); +return x_1041; } } } } else { -uint8_t x_60; -lean_dec(x_13); +uint8_t x_1111; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15024,289 +17834,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_1111 = !lean_is_exclusive(x_994); +if (x_1111 == 0) { -return x_43; +return x_994; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; -} +lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; +x_1112 = lean_ctor_get(x_994, 0); +x_1113 = lean_ctor_get(x_994, 1); +lean_inc(x_1113); +lean_inc(x_1112); +lean_dec(x_994); +x_1114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1114, 0, x_1112); +lean_ctor_set(x_1114, 1, x_1113); +return x_1114; } } } -} -else +case 10: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +lean_object* x_1115; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_11); +x_1115 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_1115) == 0) +{ +lean_object* x_1116; +x_1116 = lean_ctor_get(x_1115, 0); +lean_inc(x_1116); +if (lean_obj_tag(x_1116) == 0) +{ +lean_object* x_1117; lean_object* x_1118; +lean_dec(x_3); +lean_dec(x_1); +x_1117 = lean_ctor_get(x_1115, 1); +lean_inc(x_1117); +lean_dec(x_1115); +x_1118 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_1117); +return x_1118; +} +else +{ +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_1145; lean_object* x_1146; 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; +x_1119 = lean_ctor_get(x_1115, 1); +lean_inc(x_1119); +lean_dec(x_1115); +x_1120 = lean_ctor_get(x_1116, 0); +lean_inc(x_1120); +lean_dec(x_1116); +x_1121 = lean_st_ref_get(x_9, x_1119); +x_1122 = lean_ctor_get(x_1121, 1); +lean_inc(x_1122); +lean_dec(x_1121); +x_1123 = lean_st_ref_get(x_7, x_1122); +x_1124 = lean_ctor_get(x_1123, 0); +lean_inc(x_1124); +x_1125 = lean_ctor_get(x_1123, 1); +lean_inc(x_1125); +lean_dec(x_1123); +x_1164 = lean_st_ref_get(x_9, x_1125); +x_1165 = lean_ctor_get(x_1164, 1); +lean_inc(x_1165); +lean_dec(x_1164); +x_1166 = lean_st_ref_take(x_7, x_1165); +x_1167 = lean_ctor_get(x_1166, 0); +lean_inc(x_1167); +x_1168 = lean_ctor_get(x_1166, 1); +lean_inc(x_1168); +lean_dec(x_1166); +x_1169 = lean_ctor_get(x_1167, 0); +lean_inc(x_1169); +x_1170 = lean_ctor_get(x_1167, 2); +lean_inc(x_1170); +lean_dec(x_1167); +x_1171 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1172 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1172, 0, x_1169); +lean_ctor_set(x_1172, 1, x_1171); +lean_ctor_set(x_1172, 2, x_1170); +x_1173 = lean_st_ref_set(x_7, x_1172, x_1168); +x_1174 = lean_ctor_get(x_1173, 1); +lean_inc(x_1174); +lean_dec(x_1173); +x_1175 = l_Lean_Compiler_visitMatch(x_11, x_1120, x_7, x_8, x_9, x_1174); +x_1176 = lean_ctor_get(x_1175, 0); +lean_inc(x_1176); +x_1177 = lean_ctor_get(x_1176, 1); +lean_inc(x_1177); +x_1178 = lean_ctor_get(x_1175, 1); +lean_inc(x_1178); +lean_dec(x_1175); +x_1179 = lean_ctor_get(x_1176, 0); +lean_inc(x_1179); +lean_dec(x_1176); +x_1180 = lean_ctor_get(x_1177, 0); +lean_inc(x_1180); +x_1181 = lean_ctor_get(x_1177, 1); +lean_inc(x_1181); +lean_dec(x_1177); +lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_1182 = lean_apply_7(x_2, x_1179, x_5, x_6, x_7, x_8, x_9, x_1178); +if (lean_obj_tag(x_1182) == 0) { -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; uint8_t x_1186; +x_1183 = lean_ctor_get(x_1182, 1); +lean_inc(x_1183); +lean_dec(x_1182); +x_1184 = lean_array_get_size(x_1180); +x_1185 = lean_unsigned_to_nat(0u); +x_1186 = lean_nat_dec_lt(x_1185, x_1184); +if (x_1186 == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_1187; uint8_t x_1188; +lean_dec(x_1184); +lean_dec(x_1180); +x_1187 = lean_array_get_size(x_1181); +x_1188 = lean_nat_dec_lt(x_1185, x_1187); +if (x_1188 == 0) +{ +lean_object* x_1189; +lean_dec(x_1187); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_1189 = lean_box(0); +x_1126 = x_1189; +x_1127 = x_1183; +goto block_1144; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_1190; +x_1190 = lean_nat_dec_le(x_1187, x_1187); +if (x_1190 == 0) { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +lean_object* x_1191; +lean_dec(x_1187); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_1191 = lean_box(0); +x_1126 = x_1191; +x_1127 = x_1183; +goto block_1144; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +size_t x_1192; size_t x_1193; lean_object* x_1194; lean_object* x_1195; +x_1192 = 0; +x_1193 = lean_usize_of_nat(x_1187); +lean_dec(x_1187); +x_1194 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1195 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(x_1, x_2, x_3, x_1181, x_1192, x_1193, x_1194, x_5, x_6, x_7, x_8, x_9, x_1183); +lean_dec(x_1181); +if (lean_obj_tag(x_1195) == 0) +{ +lean_object* x_1196; lean_object* x_1197; +x_1196 = lean_ctor_get(x_1195, 0); +lean_inc(x_1196); +x_1197 = lean_ctor_get(x_1195, 1); +lean_inc(x_1197); +lean_dec(x_1195); +x_1126 = x_1196; +x_1127 = x_1197; +goto block_1144; +} +else +{ +lean_object* x_1198; lean_object* x_1199; +x_1198 = lean_ctor_get(x_1195, 0); +lean_inc(x_1198); +x_1199 = lean_ctor_get(x_1195, 1); +lean_inc(x_1199); +lean_dec(x_1195); +x_1145 = x_1198; +x_1146 = x_1199; +goto block_1163; +} } } } else { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +uint8_t x_1200; +x_1200 = lean_nat_dec_le(x_1184, x_1184); +if (x_1200 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +lean_object* x_1201; uint8_t x_1202; +lean_dec(x_1184); +lean_dec(x_1180); +x_1201 = lean_array_get_size(x_1181); +x_1202 = lean_nat_dec_lt(x_1185, x_1201); +if (x_1202 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_1203; +lean_dec(x_1201); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_1203 = lean_box(0); +x_1126 = x_1203; +x_1127 = x_1183; +goto block_1144; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +uint8_t x_1204; +x_1204 = lean_nat_dec_le(x_1201, x_1201); +if (x_1204 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_1205; +lean_dec(x_1201); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_1205 = lean_box(0); +x_1126 = x_1205; +x_1127 = x_1183; +goto block_1144; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +size_t x_1206; size_t x_1207; lean_object* x_1208; lean_object* x_1209; +x_1206 = 0; +x_1207 = lean_usize_of_nat(x_1201); +lean_dec(x_1201); +x_1208 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1209 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(x_1, x_2, x_3, x_1181, x_1206, x_1207, x_1208, x_5, x_6, x_7, x_8, x_9, x_1183); +lean_dec(x_1181); +if (lean_obj_tag(x_1209) == 0) +{ +lean_object* x_1210; lean_object* x_1211; +x_1210 = lean_ctor_get(x_1209, 0); +lean_inc(x_1210); +x_1211 = lean_ctor_get(x_1209, 1); +lean_inc(x_1211); +lean_dec(x_1209); +x_1126 = x_1210; +x_1127 = x_1211; +goto block_1144; +} +else +{ +lean_object* x_1212; lean_object* x_1213; +x_1212 = lean_ctor_get(x_1209, 0); +lean_inc(x_1212); +x_1213 = lean_ctor_get(x_1209, 1); +lean_inc(x_1213); +lean_dec(x_1209); +x_1145 = x_1212; +x_1146 = x_1213; +goto block_1163; +} } } } else { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); +size_t x_1214; size_t x_1215; lean_object* x_1216; lean_object* x_1217; +x_1214 = 0; +x_1215 = lean_usize_of_nat(x_1184); +lean_dec(x_1184); +x_1216 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +lean_inc(x_2); +x_1217 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(x_2, x_1180, x_1214, x_1215, x_1216, x_5, x_6, x_7, x_8, x_9, x_1183); +lean_dec(x_1180); +if (lean_obj_tag(x_1217) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_1218; lean_object* x_1219; uint8_t x_1220; +x_1218 = lean_ctor_get(x_1217, 1); +lean_inc(x_1218); +lean_dec(x_1217); +x_1219 = lean_array_get_size(x_1181); +x_1220 = lean_nat_dec_lt(x_1185, x_1219); +if (x_1220 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_1219); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_1126 = x_1216; +x_1127 = x_1218; +goto block_1144; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_1221; +x_1221 = lean_nat_dec_le(x_1219, x_1219); +if (x_1221 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); +lean_dec(x_1219); +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +x_1126 = x_1216; +x_1127 = x_1218; +goto block_1144; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +else +{ +size_t x_1222; lean_object* x_1223; +x_1222 = lean_usize_of_nat(x_1219); +lean_dec(x_1219); +lean_inc(x_9); +lean_inc(x_7); +x_1223 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_1, x_2, x_3, x_1181, x_1214, x_1222, x_1216, x_5, x_6, x_7, x_8, x_9, x_1218); +lean_dec(x_1181); +if (lean_obj_tag(x_1223) == 0) +{ +lean_object* x_1224; lean_object* x_1225; +x_1224 = lean_ctor_get(x_1223, 0); +lean_inc(x_1224); +x_1225 = lean_ctor_get(x_1223, 1); +lean_inc(x_1225); +lean_dec(x_1223); +x_1126 = x_1224; +x_1127 = x_1225; +goto block_1144; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_1226; lean_object* x_1227; +x_1226 = lean_ctor_get(x_1223, 0); +lean_inc(x_1226); +x_1227 = lean_ctor_get(x_1223, 1); +lean_inc(x_1227); +lean_dec(x_1223); +x_1145 = x_1226; +x_1146 = x_1227; +goto block_1163; +} } } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); +lean_object* x_1228; lean_object* x_1229; +lean_dec(x_1181); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - x_106 = lean_box(0); +x_1228 = lean_ctor_get(x_1217, 0); +lean_inc(x_1228); +x_1229 = lean_ctor_get(x_1217, 1); +lean_inc(x_1229); +lean_dec(x_1217); +x_1145 = x_1228; +x_1146 = x_1229; +goto block_1163; } -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; } } +else +{ +lean_object* x_1230; lean_object* x_1231; +lean_dec(x_1181); +lean_dec(x_1180); +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_1230 = lean_ctor_get(x_1182, 0); +lean_inc(x_1230); +x_1231 = lean_ctor_get(x_1182, 1); +lean_inc(x_1231); +lean_dec(x_1182); +x_1145 = x_1230; +x_1146 = x_1231; +goto block_1163; +} +block_1144: +{ +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; uint8_t x_1140; +x_1128 = lean_st_ref_get(x_9, x_1127); +x_1129 = lean_ctor_get(x_1128, 1); +lean_inc(x_1129); +lean_dec(x_1128); +x_1130 = lean_st_ref_get(x_7, x_1129); +x_1131 = lean_ctor_get(x_1130, 0); +lean_inc(x_1131); +x_1132 = lean_ctor_get(x_1130, 1); +lean_inc(x_1132); +lean_dec(x_1130); +x_1133 = lean_ctor_get(x_1124, 0); +lean_inc(x_1133); +x_1134 = lean_ctor_get(x_1124, 1); +lean_inc(x_1134); +lean_dec(x_1124); +x_1135 = lean_ctor_get(x_1131, 2); +lean_inc(x_1135); +lean_dec(x_1131); +x_1136 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1136, 0, x_1133); +lean_ctor_set(x_1136, 1, x_1134); +lean_ctor_set(x_1136, 2, x_1135); +x_1137 = lean_st_ref_get(x_9, x_1132); +lean_dec(x_9); +x_1138 = lean_ctor_get(x_1137, 1); +lean_inc(x_1138); +lean_dec(x_1137); +x_1139 = lean_st_ref_set(x_7, x_1136, x_1138); +lean_dec(x_7); +x_1140 = !lean_is_exclusive(x_1139); +if (x_1140 == 0) +{ +lean_object* x_1141; +x_1141 = lean_ctor_get(x_1139, 0); +lean_dec(x_1141); +lean_ctor_set(x_1139, 0, x_1126); +return x_1139; +} +else +{ +lean_object* x_1142; lean_object* x_1143; +x_1142 = lean_ctor_get(x_1139, 1); +lean_inc(x_1142); +lean_dec(x_1139); +x_1143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1143, 0, x_1126); +lean_ctor_set(x_1143, 1, x_1142); +return x_1143; +} +} +block_1163: +{ +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; uint8_t x_1159; +x_1147 = lean_st_ref_get(x_9, x_1146); +x_1148 = lean_ctor_get(x_1147, 1); +lean_inc(x_1148); +lean_dec(x_1147); +x_1149 = lean_st_ref_get(x_7, x_1148); +x_1150 = lean_ctor_get(x_1149, 0); +lean_inc(x_1150); +x_1151 = lean_ctor_get(x_1149, 1); +lean_inc(x_1151); +lean_dec(x_1149); +x_1152 = lean_ctor_get(x_1124, 0); +lean_inc(x_1152); +x_1153 = lean_ctor_get(x_1124, 1); +lean_inc(x_1153); +lean_dec(x_1124); +x_1154 = lean_ctor_get(x_1150, 2); +lean_inc(x_1154); +lean_dec(x_1150); +x_1155 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1155, 0, x_1152); +lean_ctor_set(x_1155, 1, x_1153); +lean_ctor_set(x_1155, 2, x_1154); +x_1156 = lean_st_ref_get(x_9, x_1151); +lean_dec(x_9); +x_1157 = lean_ctor_get(x_1156, 1); +lean_inc(x_1157); +lean_dec(x_1156); +x_1158 = lean_st_ref_set(x_7, x_1155, x_1157); +lean_dec(x_7); +x_1159 = !lean_is_exclusive(x_1158); +if (x_1159 == 0) +{ +lean_object* x_1160; +x_1160 = lean_ctor_get(x_1158, 0); +lean_dec(x_1160); +lean_ctor_set_tag(x_1158, 1); +lean_ctor_set(x_1158, 0, x_1145); +return x_1158; +} +else +{ +lean_object* x_1161; lean_object* x_1162; +x_1161 = lean_ctor_get(x_1158, 1); +lean_inc(x_1161); +lean_dec(x_1158); +x_1162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1162, 0, x_1145); +lean_ctor_set(x_1162, 1, x_1161); +return x_1162; +} } } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_1232; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15314,311 +18373,538 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_1232 = !lean_is_exclusive(x_1115); +if (x_1232 == 0) { -return x_14; +return x_1115; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} +lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; +x_1233 = lean_ctor_get(x_1115, 0); +x_1234 = lean_ctor_get(x_1115, 1); +lean_inc(x_1234); +lean_inc(x_1233); +lean_dec(x_1115); +x_1235 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1235, 0, x_1233); +lean_ctor_set(x_1235, 1, x_1234); +return x_1235; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___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) { -_start: +default: { -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_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); +lean_object* x_1236; +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_11); -lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); +x_1236 = l_Lean_Compiler_isCasesApp_x3f(x_11, x_8, x_9, x_10); +if (lean_obj_tag(x_1236) == 0) +{ +lean_object* x_1237; +x_1237 = lean_ctor_get(x_1236, 0); +lean_inc(x_1237); +if (lean_obj_tag(x_1237) == 0) +{ +lean_object* x_1238; lean_object* x_1239; +lean_dec(x_3); +lean_dec(x_1); +x_1238 = lean_ctor_get(x_1236, 1); +lean_inc(x_1238); +lean_dec(x_1236); +x_1239 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_1238); +return x_1239; +} +else +{ +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_1266; lean_object* x_1267; 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; +x_1240 = lean_ctor_get(x_1236, 1); +lean_inc(x_1240); +lean_dec(x_1236); +x_1241 = lean_ctor_get(x_1237, 0); +lean_inc(x_1241); +lean_dec(x_1237); +x_1242 = lean_st_ref_get(x_9, x_1240); +x_1243 = lean_ctor_get(x_1242, 1); +lean_inc(x_1243); +lean_dec(x_1242); +x_1244 = lean_st_ref_get(x_7, x_1243); +x_1245 = lean_ctor_get(x_1244, 0); +lean_inc(x_1245); +x_1246 = lean_ctor_get(x_1244, 1); +lean_inc(x_1246); +lean_dec(x_1244); +x_1285 = lean_st_ref_get(x_9, x_1246); +x_1286 = lean_ctor_get(x_1285, 1); +lean_inc(x_1286); +lean_dec(x_1285); +x_1287 = lean_st_ref_take(x_7, x_1286); +x_1288 = lean_ctor_get(x_1287, 0); +lean_inc(x_1288); +x_1289 = lean_ctor_get(x_1287, 1); +lean_inc(x_1289); +lean_dec(x_1287); +x_1290 = lean_ctor_get(x_1288, 0); +lean_inc(x_1290); +x_1291 = lean_ctor_get(x_1288, 2); +lean_inc(x_1291); +lean_dec(x_1288); +x_1292 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1293 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1293, 0, x_1290); +lean_ctor_set(x_1293, 1, x_1292); +lean_ctor_set(x_1293, 2, x_1291); +x_1294 = lean_st_ref_set(x_7, x_1293, x_1289); +x_1295 = lean_ctor_get(x_1294, 1); +lean_inc(x_1295); +lean_dec(x_1294); +x_1296 = l_Lean_Compiler_visitMatch(x_11, x_1241, x_7, x_8, x_9, x_1295); +x_1297 = lean_ctor_get(x_1296, 0); +lean_inc(x_1297); +x_1298 = lean_ctor_get(x_1297, 1); +lean_inc(x_1298); +x_1299 = lean_ctor_get(x_1296, 1); +lean_inc(x_1299); +lean_dec(x_1296); +x_1300 = lean_ctor_get(x_1297, 0); +lean_inc(x_1300); +lean_dec(x_1297); +x_1301 = lean_ctor_get(x_1298, 0); +lean_inc(x_1301); +x_1302 = lean_ctor_get(x_1298, 1); +lean_inc(x_1302); +lean_dec(x_1298); +lean_inc(x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_1303 = lean_apply_7(x_2, x_1300, x_5, x_6, x_7, x_8, x_9, x_1299); +if (lean_obj_tag(x_1303) == 0) { -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; uint8_t x_1307; +x_1304 = lean_ctor_get(x_1303, 1); +lean_inc(x_1304); +lean_dec(x_1303); +x_1305 = lean_array_get_size(x_1301); +x_1306 = lean_unsigned_to_nat(0u); +x_1307 = lean_nat_dec_lt(x_1306, x_1305); +if (x_1307 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +lean_object* x_1308; uint8_t x_1309; +lean_dec(x_1305); +lean_dec(x_1301); +x_1308 = lean_array_get_size(x_1302); +x_1309 = lean_nat_dec_lt(x_1306, x_1308); +if (x_1309 == 0) { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_1310; +lean_dec(x_1308); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_1310 = lean_box(0); +x_1247 = x_1310; +x_1248 = x_1304; +goto block_1265; } else { -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +uint8_t x_1311; +x_1311 = lean_nat_dec_le(x_1308, x_1308); +if (x_1311 == 0) { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); +lean_object* x_1312; +lean_dec(x_1308); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_1312 = lean_box(0); +x_1247 = x_1312; +x_1248 = x_1304; +goto block_1265; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; +size_t x_1313; size_t x_1314; lean_object* x_1315; lean_object* x_1316; +x_1313 = 0; +x_1314 = lean_usize_of_nat(x_1308); +lean_dec(x_1308); +x_1315 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1316 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(x_1, x_2, x_3, x_1302, x_1313, x_1314, x_1315, x_5, x_6, x_7, x_8, x_9, x_1304); +lean_dec(x_1302); +if (lean_obj_tag(x_1316) == 0) +{ +lean_object* x_1317; lean_object* x_1318; +x_1317 = lean_ctor_get(x_1316, 0); +lean_inc(x_1317); +x_1318 = lean_ctor_get(x_1316, 1); +lean_inc(x_1318); +lean_dec(x_1316); +x_1247 = x_1317; +x_1248 = x_1318; +goto block_1265; +} +else +{ +lean_object* x_1319; lean_object* x_1320; +x_1319 = lean_ctor_get(x_1316, 0); +lean_inc(x_1319); +x_1320 = lean_ctor_get(x_1316, 1); +lean_inc(x_1320); +lean_dec(x_1316); +x_1266 = x_1319; +x_1267 = x_1320; +goto block_1284; +} } } } else { -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) +uint8_t x_1321; +x_1321 = lean_nat_dec_le(x_1305, x_1305); +if (x_1321 == 0) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +lean_object* x_1322; uint8_t x_1323; +lean_dec(x_1305); +lean_dec(x_1301); +x_1322 = lean_array_get_size(x_1302); +x_1323 = lean_nat_dec_lt(x_1306, x_1322); +if (x_1323 == 0) { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_1324; +lean_dec(x_1322); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +x_1324 = lean_box(0); +x_1247 = x_1324; +x_1248 = x_1304; +goto block_1265; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +uint8_t x_1325; +x_1325 = lean_nat_dec_le(x_1322, x_1322); +if (x_1325 == 0) { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); +lean_object* x_1326; +lean_dec(x_1322); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); 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); -lean_ctor_set(x_14, 0, x_35); -return x_14; +x_1326 = lean_box(0); +x_1247 = x_1326; +x_1248 = x_1304; +goto block_1265; } else { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; +size_t x_1327; size_t x_1328; lean_object* x_1329; lean_object* x_1330; +x_1327 = 0; +x_1328 = lean_usize_of_nat(x_1322); +lean_dec(x_1322); +x_1329 = lean_box(0); +lean_inc(x_9); +lean_inc(x_7); +x_1330 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(x_1, x_2, x_3, x_1302, x_1327, x_1328, x_1329, x_5, x_6, x_7, x_8, x_9, x_1304); +lean_dec(x_1302); +if (lean_obj_tag(x_1330) == 0) +{ +lean_object* x_1331; lean_object* x_1332; +x_1331 = lean_ctor_get(x_1330, 0); +lean_inc(x_1331); +x_1332 = lean_ctor_get(x_1330, 1); +lean_inc(x_1332); +lean_dec(x_1330); +x_1247 = x_1331; +x_1248 = x_1332; +goto block_1265; +} +else +{ +lean_object* x_1333; lean_object* x_1334; +x_1333 = lean_ctor_get(x_1330, 0); +lean_inc(x_1333); +x_1334 = lean_ctor_get(x_1330, 1); +lean_inc(x_1334); +lean_dec(x_1330); +x_1266 = x_1333; +x_1267 = x_1334; +goto block_1284; +} } } } else { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); +size_t x_1335; size_t x_1336; lean_object* x_1337; lean_object* x_1338; +x_1335 = 0; +x_1336 = lean_usize_of_nat(x_1305); +lean_dec(x_1305); +x_1337 = lean_box(0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_inc(x_2); +x_1338 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(x_2, x_1301, x_1335, x_1336, x_1337, x_5, x_6, x_7, x_8, x_9, x_1304); +lean_dec(x_1301); +if (lean_obj_tag(x_1338) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +lean_object* x_1339; lean_object* x_1340; uint8_t x_1341; +x_1339 = lean_ctor_get(x_1338, 1); +lean_inc(x_1339); +lean_dec(x_1338); +x_1340 = lean_array_get_size(x_1302); +x_1341 = lean_nat_dec_lt(x_1306, x_1340); +if (x_1341 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_1340); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_1247 = x_1337; +x_1248 = x_1339; +goto block_1265; } else { -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +uint8_t x_1342; +x_1342 = lean_nat_dec_le(x_1340, x_1340); +if (x_1342 == 0) { -lean_dec(x_47); -lean_dec(x_13); +lean_dec(x_1340); +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_1247 = x_1337; +x_1248 = x_1339; +goto block_1265; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; +size_t x_1343; lean_object* x_1344; +x_1343 = lean_usize_of_nat(x_1340); +lean_dec(x_1340); +lean_inc(x_9); +lean_inc(x_7); +x_1344 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_1, x_2, x_3, x_1302, x_1335, x_1343, x_1337, x_5, x_6, x_7, x_8, x_9, x_1339); +lean_dec(x_1302); +if (lean_obj_tag(x_1344) == 0) +{ +lean_object* x_1345; lean_object* x_1346; +x_1345 = lean_ctor_get(x_1344, 0); +lean_inc(x_1345); +x_1346 = lean_ctor_get(x_1344, 1); +lean_inc(x_1346); +lean_dec(x_1344); +x_1247 = x_1345; +x_1248 = x_1346; +goto block_1265; +} +else +{ +lean_object* x_1347; lean_object* x_1348; +x_1347 = lean_ctor_get(x_1344, 0); +lean_inc(x_1347); +x_1348 = lean_ctor_get(x_1344, 1); +lean_inc(x_1348); +lean_dec(x_1344); +x_1266 = x_1347; +x_1267 = x_1348; +goto block_1284; +} } } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_1349; lean_object* x_1350; +lean_dec(x_1302); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_1349 = lean_ctor_get(x_1338, 0); +lean_inc(x_1349); +x_1350 = lean_ctor_get(x_1338, 1); +lean_inc(x_1350); +lean_dec(x_1338); +x_1266 = x_1349; +x_1267 = x_1350; +goto block_1284; +} +} +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); +lean_object* x_1351; lean_object* x_1352; +lean_dec(x_1302); +lean_dec(x_1301); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_1351 = lean_ctor_get(x_1303, 0); +lean_inc(x_1351); +x_1352 = lean_ctor_get(x_1303, 1); +lean_inc(x_1352); +lean_dec(x_1303); +x_1266 = x_1351; +x_1267 = x_1352; +goto block_1284; +} +block_1265: +{ +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; uint8_t x_1261; +x_1249 = lean_st_ref_get(x_9, x_1248); +x_1250 = lean_ctor_get(x_1249, 1); +lean_inc(x_1250); +lean_dec(x_1249); +x_1251 = lean_st_ref_get(x_7, x_1250); +x_1252 = lean_ctor_get(x_1251, 0); +lean_inc(x_1252); +x_1253 = lean_ctor_get(x_1251, 1); +lean_inc(x_1253); +lean_dec(x_1251); +x_1254 = lean_ctor_get(x_1245, 0); +lean_inc(x_1254); +x_1255 = lean_ctor_get(x_1245, 1); +lean_inc(x_1255); +lean_dec(x_1245); +x_1256 = lean_ctor_get(x_1252, 2); +lean_inc(x_1256); +lean_dec(x_1252); +x_1257 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1257, 0, x_1254); +lean_ctor_set(x_1257, 1, x_1255); +lean_ctor_set(x_1257, 2, x_1256); +x_1258 = lean_st_ref_get(x_9, x_1253); +lean_dec(x_9); +x_1259 = lean_ctor_get(x_1258, 1); +lean_inc(x_1259); +lean_dec(x_1258); +x_1260 = lean_st_ref_set(x_7, x_1257, x_1259); +lean_dec(x_7); +x_1261 = !lean_is_exclusive(x_1260); +if (x_1261 == 0) +{ +lean_object* x_1262; +x_1262 = lean_ctor_get(x_1260, 0); +lean_dec(x_1262); +lean_ctor_set(x_1260, 0, x_1247); +return x_1260; +} +else +{ +lean_object* x_1263; lean_object* x_1264; +x_1263 = lean_ctor_get(x_1260, 1); +lean_inc(x_1263); +lean_dec(x_1260); +x_1264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1264, 0, x_1247); +lean_ctor_set(x_1264, 1, x_1263); +return x_1264; +} +} +block_1284: +{ +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; uint8_t x_1280; +x_1268 = lean_st_ref_get(x_9, x_1267); +x_1269 = lean_ctor_get(x_1268, 1); +lean_inc(x_1269); +lean_dec(x_1268); +x_1270 = lean_st_ref_get(x_7, x_1269); +x_1271 = lean_ctor_get(x_1270, 0); +lean_inc(x_1271); +x_1272 = lean_ctor_get(x_1270, 1); +lean_inc(x_1272); +lean_dec(x_1270); +x_1273 = lean_ctor_get(x_1245, 0); +lean_inc(x_1273); +x_1274 = lean_ctor_get(x_1245, 1); +lean_inc(x_1274); +lean_dec(x_1245); +x_1275 = lean_ctor_get(x_1271, 2); +lean_inc(x_1275); +lean_dec(x_1271); +x_1276 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1276, 0, x_1273); +lean_ctor_set(x_1276, 1, x_1274); +lean_ctor_set(x_1276, 2, x_1275); +x_1277 = lean_st_ref_get(x_9, x_1272); +lean_dec(x_9); +x_1278 = lean_ctor_get(x_1277, 1); +lean_inc(x_1278); +lean_dec(x_1277); +x_1279 = lean_st_ref_set(x_7, x_1276, x_1278); +lean_dec(x_7); +x_1280 = !lean_is_exclusive(x_1279); +if (x_1280 == 0) +{ +lean_object* x_1281; +x_1281 = lean_ctor_get(x_1279, 0); +lean_dec(x_1281); +lean_ctor_set_tag(x_1279, 1); +lean_ctor_set(x_1279, 0, x_1266); +return x_1279; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +lean_object* x_1282; lean_object* x_1283; +x_1282 = lean_ctor_get(x_1279, 1); +lean_inc(x_1282); +lean_dec(x_1279); +x_1283 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1283, 0, x_1266); +lean_ctor_set(x_1283, 1, x_1282); +return x_1283; } } } } else { -uint8_t x_60; -lean_dec(x_13); +uint8_t x_1353; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15626,135 +18912,300 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) +x_1353 = !lean_is_exclusive(x_1236); +if (x_1353 == 0) { -return x_43; +return x_1236; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; +lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; +x_1354 = lean_ctor_get(x_1236, 0); +x_1355 = lean_ctor_get(x_1236, 1); +lean_inc(x_1355); +lean_inc(x_1354); +lean_dec(x_1236); +x_1356 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1356, 0, x_1354); +lean_ctor_set(x_1356, 1, x_1355); +return x_1356; } } } } } -else -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +case 6: { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_1357; lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_1357 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +return x_1357; +} +case 8: +{ +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; +x_1358 = lean_st_ref_get(x_9, x_10); +x_1359 = lean_ctor_get(x_1358, 1); +lean_inc(x_1359); +lean_dec(x_1358); +x_1360 = lean_st_ref_get(x_7, x_1359); +x_1361 = lean_ctor_get(x_1360, 0); +lean_inc(x_1361); +x_1362 = lean_ctor_get(x_1360, 1); +lean_inc(x_1362); +lean_dec(x_1360); +x_1363 = lean_st_ref_get(x_9, x_1362); +x_1364 = lean_ctor_get(x_1363, 1); +lean_inc(x_1364); +lean_dec(x_1363); +x_1365 = lean_st_ref_take(x_7, x_1364); +x_1366 = lean_ctor_get(x_1365, 0); +lean_inc(x_1366); +x_1367 = lean_ctor_get(x_1365, 1); +lean_inc(x_1367); +lean_dec(x_1365); +x_1368 = lean_ctor_get(x_1366, 0); +lean_inc(x_1368); +x_1369 = lean_ctor_get(x_1366, 2); +lean_inc(x_1369); +lean_dec(x_1366); +x_1370 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1371 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1371, 0, x_1368); +lean_ctor_set(x_1371, 1, x_1370); +lean_ctor_set(x_1371, 2, x_1369); +x_1372 = lean_st_ref_set(x_7, x_1371, x_1367); +x_1373 = lean_ctor_get(x_1372, 1); +lean_inc(x_1373); +lean_dec(x_1372); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_3); +x_1374 = l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_3, x_5, x_6, x_11, x_1370, x_7, x_8, x_9, x_1373); +if (lean_obj_tag(x_1374) == 0) +{ +lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; +x_1375 = lean_ctor_get(x_1374, 0); +lean_inc(x_1375); +x_1376 = lean_ctor_get(x_1374, 1); +lean_inc(x_1376); +lean_dec(x_1374); +lean_inc(x_9); +lean_inc(x_7); +x_1377 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_1375, x_5, x_6, x_7, x_8, x_9, x_1376); +lean_dec(x_1375); +if (lean_obj_tag(x_1377) == 0) +{ +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; uint8_t x_1392; +x_1378 = lean_ctor_get(x_1377, 0); +lean_inc(x_1378); +x_1379 = lean_ctor_get(x_1377, 1); +lean_inc(x_1379); +lean_dec(x_1377); +x_1380 = lean_st_ref_get(x_9, x_1379); +x_1381 = lean_ctor_get(x_1380, 1); +lean_inc(x_1381); +lean_dec(x_1380); +x_1382 = lean_st_ref_get(x_7, x_1381); +x_1383 = lean_ctor_get(x_1382, 0); +lean_inc(x_1383); +x_1384 = lean_ctor_get(x_1382, 1); +lean_inc(x_1384); +lean_dec(x_1382); +x_1385 = lean_ctor_get(x_1361, 0); +lean_inc(x_1385); +x_1386 = lean_ctor_get(x_1361, 1); +lean_inc(x_1386); +lean_dec(x_1361); +x_1387 = lean_ctor_get(x_1383, 2); +lean_inc(x_1387); +lean_dec(x_1383); +x_1388 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1388, 0, x_1385); +lean_ctor_set(x_1388, 1, x_1386); +lean_ctor_set(x_1388, 2, x_1387); +x_1389 = lean_st_ref_get(x_9, x_1384); +lean_dec(x_9); +x_1390 = lean_ctor_get(x_1389, 1); +lean_inc(x_1390); +lean_dec(x_1389); +x_1391 = lean_st_ref_set(x_7, x_1388, x_1390); +lean_dec(x_7); +x_1392 = !lean_is_exclusive(x_1391); +if (x_1392 == 0) +{ +lean_object* x_1393; +x_1393 = lean_ctor_get(x_1391, 0); +lean_dec(x_1393); +lean_ctor_set(x_1391, 0, x_1378); +return x_1391; +} +else +{ +lean_object* x_1394; lean_object* x_1395; +x_1394 = lean_ctor_get(x_1391, 1); +lean_inc(x_1394); +lean_dec(x_1391); +x_1395 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1395, 0, x_1378); +lean_ctor_set(x_1395, 1, x_1394); +return x_1395; +} +} +else +{ +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; uint8_t x_1410; +x_1396 = lean_ctor_get(x_1377, 0); +lean_inc(x_1396); +x_1397 = lean_ctor_get(x_1377, 1); +lean_inc(x_1397); +lean_dec(x_1377); +x_1398 = lean_st_ref_get(x_9, x_1397); +x_1399 = lean_ctor_get(x_1398, 1); +lean_inc(x_1399); +lean_dec(x_1398); +x_1400 = lean_st_ref_get(x_7, x_1399); +x_1401 = lean_ctor_get(x_1400, 0); +lean_inc(x_1401); +x_1402 = lean_ctor_get(x_1400, 1); +lean_inc(x_1402); +lean_dec(x_1400); +x_1403 = lean_ctor_get(x_1361, 0); +lean_inc(x_1403); +x_1404 = lean_ctor_get(x_1361, 1); +lean_inc(x_1404); +lean_dec(x_1361); +x_1405 = lean_ctor_get(x_1401, 2); +lean_inc(x_1405); +lean_dec(x_1401); +x_1406 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1406, 0, x_1403); +lean_ctor_set(x_1406, 1, x_1404); +lean_ctor_set(x_1406, 2, x_1405); +x_1407 = lean_st_ref_get(x_9, x_1402); +lean_dec(x_9); +x_1408 = lean_ctor_get(x_1407, 1); +lean_inc(x_1408); +lean_dec(x_1407); +x_1409 = lean_st_ref_set(x_7, x_1406, x_1408); +lean_dec(x_7); +x_1410 = !lean_is_exclusive(x_1409); +if (x_1410 == 0) +{ +lean_object* x_1411; +x_1411 = lean_ctor_get(x_1409, 0); +lean_dec(x_1411); +lean_ctor_set_tag(x_1409, 1); +lean_ctor_set(x_1409, 0, x_1396); +return x_1409; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +lean_object* x_1412; lean_object* x_1413; +x_1412 = lean_ctor_get(x_1409, 1); +lean_inc(x_1412); +lean_dec(x_1409); +x_1413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1413, 0, x_1396); +lean_ctor_set(x_1413, 1, x_1412); +return x_1413; +} +} +} +else { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); +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; uint8_t x_1428; lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +x_1414 = lean_ctor_get(x_1374, 0); +lean_inc(x_1414); +x_1415 = lean_ctor_get(x_1374, 1); +lean_inc(x_1415); +lean_dec(x_1374); +x_1416 = lean_st_ref_get(x_9, x_1415); +x_1417 = lean_ctor_get(x_1416, 1); +lean_inc(x_1417); +lean_dec(x_1416); +x_1418 = lean_st_ref_get(x_7, x_1417); +x_1419 = lean_ctor_get(x_1418, 0); +lean_inc(x_1419); +x_1420 = lean_ctor_get(x_1418, 1); +lean_inc(x_1420); +lean_dec(x_1418); +x_1421 = lean_ctor_get(x_1361, 0); +lean_inc(x_1421); +x_1422 = lean_ctor_get(x_1361, 1); +lean_inc(x_1422); +lean_dec(x_1361); +x_1423 = lean_ctor_get(x_1419, 2); +lean_inc(x_1423); +lean_dec(x_1419); +x_1424 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1424, 0, x_1421); +lean_ctor_set(x_1424, 1, x_1422); +lean_ctor_set(x_1424, 2, x_1423); +x_1425 = lean_st_ref_get(x_9, x_1420); +lean_dec(x_9); +x_1426 = lean_ctor_get(x_1425, 1); +lean_inc(x_1426); +lean_dec(x_1425); +x_1427 = lean_st_ref_set(x_7, x_1424, x_1426); +lean_dec(x_7); +x_1428 = !lean_is_exclusive(x_1427); +if (x_1428 == 0) +{ +lean_object* x_1429; +x_1429 = lean_ctor_get(x_1427, 0); +lean_dec(x_1429); +lean_ctor_set_tag(x_1427, 1); +lean_ctor_set(x_1427, 0, x_1414); +return x_1427; } else { -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +lean_object* x_1430; lean_object* x_1431; +x_1430 = lean_ctor_get(x_1427, 1); +lean_inc(x_1430); +lean_dec(x_1427); +x_1431 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1431, 0, x_1414); +lean_ctor_set(x_1431, 1, x_1430); +return x_1431; } } } -else +case 10: { -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_1432; lean_object* x_1433; +lean_dec(x_11); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_83; +x_1432 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_1433 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_1432, x_5, x_6, x_7, x_8, x_9, x_10); +return x_1433; } -else +case 11: { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +lean_object* x_1434; +lean_dec(x_3); +lean_dec(x_1); +x_1434 = lean_apply_7(x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +return x_1434; +} +default: { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); +lean_object* x_1435; lean_object* x_1436; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15762,14863 +19213,10384 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_64); -return x_86; +x_1435 = lean_box(0); +x_1436 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1436, 0, x_1435); +lean_ctor_set(x_1436, 1, x_10); +return x_1436; } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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: { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +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; } } +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___boxed), 8, 0); +return x_1; +} } -else +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2() { +_start: { -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn), 7, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue(lean_object* x_1, lean_object* x_2, 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_2) == 6) +{ +lean_object* x_9; 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_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; +x_19 = lean_st_ref_get(x_7, x_8); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_st_ref_get(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_43 = lean_st_ref_get(x_7, x_23); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_st_ref_take(x_5, x_44); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_46, 2); +lean_inc(x_49); +lean_dec(x_46); +x_50 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_48); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = lean_st_ref_set(x_5, x_51, x_47); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_inc(x_2); +x_54 = l_Lean_Compiler_visitLambda(x_2, x_5, x_6, x_7, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_55, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_array_get_size(x_57); +lean_dec(x_57); +lean_inc(x_1); +x_60 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_addCandidate(x_1, x_59, x_3, x_4, x_5, x_6, x_7, x_56); +lean_dec(x_3); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_1); +x_63 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1; +x_64 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2; +x_65 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3; lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +x_66 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_63, x_64, x_65, x_58, x_62, x_4, x_5, x_6, x_7, x_61); +lean_dec(x_58); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); +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; uint8_t x_81; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_st_ref_get(x_7, x_68); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = lean_st_ref_get(x_5, 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); +lean_dec(x_71); +x_74 = lean_ctor_get(x_22, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_22, 1); +lean_inc(x_75); +lean_dec(x_22); +x_76 = lean_ctor_get(x_72, 2); +lean_inc(x_76); +lean_dec(x_72); +x_77 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_75); +lean_ctor_set(x_77, 2, x_76); +x_78 = lean_st_ref_get(x_7, x_73); +lean_dec(x_7); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_st_ref_set(x_5, x_77, x_79); +lean_dec(x_5); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; +x_82 = lean_ctor_get(x_80, 0); +lean_dec(x_82); +lean_ctor_set(x_80, 0, x_67); +x_9 = x_80; +goto block_18; } -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +else { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +lean_dec(x_80); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_67); +lean_ctor_set(x_84, 1, x_83); +x_9 = x_84; +goto block_18; } -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) -{ -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +lean_object* x_85; lean_object* x_86; +x_85 = lean_ctor_get(x_66, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_66, 1); +lean_inc(x_86); +lean_dec(x_66); +x_24 = x_85; +x_25 = x_86; +goto block_42; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; +block_18: +{ +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; +x_11 = lean_ctor_get(x_9, 0); +lean_dec(x_11); +lean_ctor_set(x_9, 0, x_2); +return x_9; } else { -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_12); +return x_13; } } +else +{ +uint8_t x_14; +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) +{ +return x_9; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +block_42: +{ +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; +x_26 = lean_st_ref_get(x_7, x_25); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_get(x_5, x_27); +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_22, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_dec(x_22); +x_33 = lean_ctor_get(x_29, 2); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 3, 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); +x_35 = lean_st_ref_get(x_7, x_30); lean_dec(x_7); -lean_dec(x_6); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_set(x_5, x_34, x_36); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - x_106 = lean_box(0); +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_tag(x_37, 1); +lean_ctor_set(x_37, 0, x_24); +x_9 = x_37; +goto block_18; } -if (lean_is_scalar(x_106)) { - x_107 = lean_alloc_ctor(1, 2, 0); -} else { - x_107 = x_106; +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(1, 2, 0); +lean_ctor_set(x_41, 0, x_24); +lean_ctor_set(x_41, 1, x_40); +x_9 = x_41; +goto block_18; } -lean_ctor_set(x_107, 0, x_104); -lean_ctor_set(x_107, 1, x_105); -return x_107; } } +else +{ +lean_object* x_87; lean_object* x_88; +lean_dec(x_1); +x_87 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1; +lean_inc(x_2); +x_88 = l_Lean_Compiler_JoinPoints_forEachFVar___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__3(x_2, x_87, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_88) == 0) +{ +uint8_t x_89; +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) +{ +lean_object* x_90; +x_90 = lean_ctor_get(x_88, 0); +lean_dec(x_90); +lean_ctor_set(x_88, 0, x_2); +return x_88; } +else +{ +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_88, 1); +lean_inc(x_91); +lean_dec(x_88); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_2); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } else { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); +uint8_t x_93; lean_dec(x_2); -lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +x_93 = !lean_is_exclusive(x_88); +if (x_93 == 0) { -return x_14; +return x_88; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_88, 0); +x_95 = lean_ctor_get(x_88, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_88); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; +} } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_13; lean_object* x_14; -x_10 = lean_ctor_get(x_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -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; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); -return x_14; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__3(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_25; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(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_2); -lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); return x_14; } -else -{ -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_13) { +_start: { -lean_object* x_33; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: +{ +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(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_2); -lean_dec(x_1); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); return x_14; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_13) { +_start: { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_11) { +_start: +{ +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(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_2); -lean_dec(x_1); -x_35 = lean_box(0); -lean_ctor_set(x_14, 0, x_35); return x_14; } -else -{ -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_42 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -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; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: +{ +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(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_2); -lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +return x_14; } -else -{ -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_7, 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_57; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_3); lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(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_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +return x_14; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: { -uint8_t x_60; -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__23(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: +{ +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(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_2); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) -{ -return x_43; +return x_14; } -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { +_start: { -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(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_2); -lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +return x_14; } -else -{ -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -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_64); -return x_83; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_3); lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(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_2); -lean_dec(x_1); -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_64); -return x_86; +return x_14; } -else -{ -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else -{ -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; +} } -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_99; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; } -else -{ -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_101; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(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_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; +return x_14; } -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; -} +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_108; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(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_2); -lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) -{ return x_14; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_13; lean_object* x_14; -x_10 = lean_ctor_get(x_4, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_14 = lean_apply_6(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -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_14, 1); -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_12); -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; uint8_t x_22; -lean_dec(x_18); -lean_dec(x_12); -x_21 = lean_array_get_size(x_13); -x_22 = lean_nat_dec_lt(x_19, x_21); -if (x_22 == 0) +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_23; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(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_2); -lean_dec(x_1); -x_23 = lean_box(0); -lean_ctor_set(x_14, 0, x_23); return x_14; } -else -{ -uint8_t x_24; -x_24 = lean_nat_dec_le(x_21, x_21); -if (x_24 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_25; -lean_dec(x_21); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_box(0); -lean_ctor_set(x_14, 0, x_25); -return x_14; +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_16; } -else -{ -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_14); -x_26 = 0; -x_27 = lean_usize_of_nat(x_21); -lean_dec(x_21); -x_28 = lean_box(0); -x_29 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(x_2, x_1, x_3, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_29; } +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_4); +return x_11; } } -else -{ -uint8_t x_30; -x_30 = lean_nat_dec_le(x_18, x_18); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -lean_dec(x_18); -lean_dec(x_12); -x_31 = lean_array_get_size(x_13); -x_32 = lean_nat_dec_lt(x_19, x_31); -if (x_32 == 0) +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_33; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); +lean_object* x_9; +x_9 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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); -x_33 = lean_box(0); -lean_ctor_set(x_14, 0, x_33); -return x_14; +return x_9; } -else -{ -uint8_t x_34; -x_34 = lean_nat_dec_le(x_31, x_31); -if (x_34 == 0) -{ -lean_object* x_35; -lean_dec(x_31); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -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); -lean_ctor_set(x_14, 0, x_35); -return x_14; } -else +LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1(lean_object* x_1) { +_start: { -size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_36 = 0; -x_37 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_38 = lean_box(0); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(x_2, x_1, x_3, x_13, x_36, x_37, x_38, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_13); -return x_39; -} +lean_object* x_2; +x_2 = l_Std_mkHashMapImp___rarg(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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: { -size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_14); -x_40 = 0; -x_41 = lean_usize_of_nat(x_18); +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_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_7, 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_st_ref_get(x_9, x_15); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_take(x_7, 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); -x_42 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(x_1, x_12, x_40, x_41, x_42, x_5, x_6, x_7, x_8, x_16); -lean_dec(x_12); -if (lean_obj_tag(x_43) == 0) +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 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; lean_object* x_30; +x_22 = lean_ctor_get(x_19, 1); +lean_dec(x_22); +x_23 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +lean_ctor_set(x_19, 1, x_23); +x_24 = lean_st_ref_set(x_7, x_19, x_20); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = l_Lean_Compiler_visitLambda(x_4, x_7, x_8, x_9, x_25); +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(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +lean_inc(x_9); +lean_inc(x_7); +x_30 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_29, x_5, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_43, 0); -lean_dec(x_46); -x_47 = lean_array_get_size(x_13); -x_48 = lean_nat_dec_lt(x_19, x_47); -if (x_48 == 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; lean_object* x_39; uint8_t x_40; +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_st_ref_get(x_9, x_32); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = lean_st_ref_get(x_7, x_34); +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(x_14, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_14, 1); +lean_inc(x_39); +lean_dec(x_14); +x_40 = !lean_is_exclusive(x_36); +if (x_40 == 0) { -lean_dec(x_47); -lean_dec(x_13); -lean_dec(x_8); +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_41 = lean_ctor_get(x_36, 1); +lean_dec(x_41); +x_42 = lean_ctor_get(x_36, 0); +lean_dec(x_42); +lean_ctor_set(x_36, 1, x_39); +lean_ctor_set(x_36, 0, x_38); +x_43 = lean_st_ref_get(x_9, x_37); +lean_dec(x_9); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_st_ref_set(x_7, x_36, x_44); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; -} -else -{ -uint8_t x_49; -x_49 = lean_nat_dec_le(x_47, x_47); -if (x_49 == 0) +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_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_43, 0, x_42); -return x_43; +lean_ctor_set(x_45, 0, x_31); +return x_45; } else { -size_t x_50; lean_object* x_51; -lean_free_object(x_43); -x_50 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_2, x_1, x_3, x_13, x_40, x_50, x_42, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_13); -return x_51; -} +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_31); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } else { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_dec(x_43); -x_53 = lean_array_get_size(x_13); -x_54 = lean_nat_dec_lt(x_19, x_53); -if (x_54 == 0) -{ -lean_object* x_55; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); +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_50 = lean_ctor_get(x_36, 2); +lean_inc(x_50); +lean_dec(x_36); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_38); +lean_ctor_set(x_51, 1, x_39); +lean_ctor_set(x_51, 2, x_50); +x_52 = lean_st_ref_get(x_9, x_37); +lean_dec(x_9); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_st_ref_set(x_7, x_51, x_53); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_42); -lean_ctor_set(x_55, 1, x_52); -return x_55; +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_54); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_31); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} } else { -uint8_t x_56; -x_56 = lean_nat_dec_le(x_53, x_53); -if (x_56 == 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_object* x_66; uint8_t x_67; +x_58 = lean_ctor_get(x_30, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_30, 1); +lean_inc(x_59); +lean_dec(x_30); +x_60 = lean_st_ref_get(x_9, x_59); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_st_ref_get(x_7, x_61); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_14, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_14, 1); +lean_inc(x_66); +lean_dec(x_14); +x_67 = !lean_is_exclusive(x_63); +if (x_67 == 0) { -lean_object* x_57; -lean_dec(x_53); -lean_dec(x_13); -lean_dec(x_8); +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_63, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_63, 0); +lean_dec(x_69); +lean_ctor_set(x_63, 1, x_66); +lean_ctor_set(x_63, 0, x_65); +x_70 = lean_st_ref_get(x_9, x_64); +lean_dec(x_9); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = lean_st_ref_set(x_7, x_63, x_71); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_42); -lean_ctor_set(x_57, 1, x_52); -return x_57; +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) +{ +lean_object* x_74; +x_74 = lean_ctor_get(x_72, 0); +lean_dec(x_74); +lean_ctor_set_tag(x_72, 1); +lean_ctor_set(x_72, 0, x_58); +return x_72; } else { -size_t x_58; lean_object* x_59; -x_58 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_59 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_2, x_1, x_3, x_13, x_40, x_58, x_42, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_13); -return x_59; -} -} +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +lean_dec(x_72); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_58); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } else { -uint8_t x_60; -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_43); -if (x_60 == 0) -{ -return x_43; +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_63, 2); +lean_inc(x_77); +lean_dec(x_63); +x_78 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_78, 0, x_65); +lean_ctor_set(x_78, 1, x_66); +lean_ctor_set(x_78, 2, x_77); +x_79 = lean_st_ref_get(x_9, x_64); +lean_dec(x_9); +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = lean_st_ref_set(x_7, x_78, x_80); +lean_dec(x_7); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); +} +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(1, 2, 0); +} else { + x_84 = x_83; + lean_ctor_set_tag(x_84, 1); +} +lean_ctor_set(x_84, 0, x_58); +lean_ctor_set(x_84, 1, x_82); +return x_84; +} +} } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_43, 0); -x_62 = lean_ctor_get(x_43, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_43); -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; +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_85 = lean_ctor_get(x_19, 0); +x_86 = lean_ctor_get(x_19, 2); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_19); +x_87 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_88 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_88, 2, x_86); +x_89 = lean_st_ref_set(x_7, x_88, x_20); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); +x_91 = l_Lean_Compiler_visitLambda(x_4, x_7, x_8, x_9, x_90); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +lean_inc(x_9); +lean_inc(x_7); +x_95 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_94, x_5, x_6, x_7, x_8, x_9, x_93); +lean_dec(x_94); +if (lean_obj_tag(x_95) == 0) +{ +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; +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_st_ref_get(x_9, x_97); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = lean_st_ref_get(x_7, x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = lean_ctor_get(x_14, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_14, 1); +lean_inc(x_104); +lean_dec(x_14); +x_105 = lean_ctor_get(x_101, 2); +lean_inc(x_105); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + lean_ctor_release(x_101, 2); + x_106 = x_101; +} else { + lean_dec_ref(x_101); + x_106 = lean_box(0); } +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(0, 3, 0); +} else { + x_107 = x_106; } +lean_ctor_set(x_107, 0, x_103); +lean_ctor_set(x_107, 1, x_104); +lean_ctor_set(x_107, 2, x_105); +x_108 = lean_st_ref_get(x_9, x_102); +lean_dec(x_9); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec(x_108); +x_110 = lean_st_ref_set(x_7, x_107, x_109); +lean_dec(x_7); +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_110); + x_112 = lean_box(0); } +if (lean_is_scalar(x_112)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_112; } +lean_ctor_set(x_113, 0, x_96); +lean_ctor_set(x_113, 1, x_111); +return x_113; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); +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; +x_114 = lean_ctor_get(x_95, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_95, 1); +lean_inc(x_115); +lean_dec(x_95); +x_116 = lean_st_ref_get(x_9, x_115); +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +x_118 = lean_st_ref_get(x_7, x_117); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_ctor_get(x_14, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_14, 1); +lean_inc(x_122); lean_dec(x_14); -x_65 = lean_array_get_size(x_12); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_65); -lean_dec(x_12); -x_68 = lean_array_get_size(x_13); -x_69 = lean_nat_dec_lt(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_64); -return x_71; +x_123 = lean_ctor_get(x_119, 2); +lean_inc(x_123); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + lean_ctor_release(x_119, 2); + x_124 = x_119; +} else { + lean_dec_ref(x_119); + x_124 = lean_box(0); } -else -{ -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; -lean_dec(x_68); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_64); -return x_74; +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(0, 3, 0); +} else { + x_125 = x_124; } -else -{ -size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; -x_75 = 0; -x_76 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_77 = lean_box(0); -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(x_2, x_1, x_3, x_13, x_75, x_76, x_77, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_78; +lean_ctor_set(x_125, 0, x_121); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_st_ref_get(x_9, x_120); +lean_dec(x_9); +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_st_ref_set(x_7, x_125, x_127); +lean_dec(x_7); +x_129 = lean_ctor_get(x_128, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_130 = x_128; +} else { + lean_dec_ref(x_128); + x_130 = lean_box(0); +} +if (lean_is_scalar(x_130)) { + x_131 = lean_alloc_ctor(1, 2, 0); +} else { + x_131 = x_130; + lean_ctor_set_tag(x_131, 1); } +lean_ctor_set(x_131, 0, x_114); +lean_ctor_set(x_131, 1, x_129); +return x_131; } } -else -{ -uint8_t x_79; -x_79 = lean_nat_dec_le(x_65, x_65); -if (x_79 == 0) +} +} +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_80; uint8_t x_81; -lean_dec(x_65); -lean_dec(x_12); -x_80 = lean_array_get_size(x_13); -x_81 = lean_nat_dec_lt(x_66, x_80); -if (x_81 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -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_64); -return x_83; +return x_1; } else { -uint8_t x_84; -x_84 = lean_nat_dec_le(x_80, x_80); -if (x_84 == 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_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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_85; lean_object* x_86; -lean_dec(x_80); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); +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_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(x_4, x_6); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -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_64); -return x_86; +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +x_4 = x_7; +goto _start; } else { -size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_87 = 0; -x_88 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_89 = lean_box(0); -x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(x_2, x_1, x_3, x_13, x_87, x_88, x_89, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_13); -return x_90; +return x_4; } } } -else -{ -size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; -x_91 = 0; -x_92 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_93 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_94 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(x_1, x_12, x_91, x_92, x_93, x_5, x_6, x_7, x_8, x_64); -lean_dec(x_12); -if (lean_obj_tag(x_94) == 0) +LEAN_EXPORT lean_object* l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(lean_object* x_1) { +_start: { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -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_94); - x_96 = lean_box(0); -} -x_97 = lean_array_get_size(x_13); -x_98 = lean_nat_dec_lt(x_66, x_97); -if (x_98 == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +lean_dec(x_1); +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) { -lean_object* x_99; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_6; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_95); -return x_99; +x_6 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +return x_6; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_97, x_97); -if (x_100 == 0) +uint8_t x_7; +x_7 = lean_nat_dec_le(x_3, x_3); +if (x_7 == 0) { -lean_object* x_101; -lean_dec(x_97); -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_8; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_96)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_96; -} -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_95); -return x_101; -} -else -{ -size_t x_102; lean_object* x_103; -lean_dec(x_96); -x_102 = lean_usize_of_nat(x_97); -lean_dec(x_97); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_2, x_1, x_3, x_13, x_91, x_102, x_93, x_5, x_6, x_7, x_8, x_95); -lean_dec(x_13); -return x_103; -} -} +x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +return x_8; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = 0; +x_10 = lean_usize_of_nat(x_3); lean_dec(x_3); +x_11 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(x_2, x_9, x_10, x_11); lean_dec(x_2); -lean_dec(x_1); -x_104 = lean_ctor_get(x_94, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_94, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_106 = x_94; -} else { - lean_dec_ref(x_94); - 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; -} +return x_12; } } } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -uint8_t x_108; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_108 = !lean_is_exclusive(x_14); -if (x_108 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -return x_14; +return x_3; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_14, 0); -x_110 = lean_ctor_get(x_14, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_14); -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; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__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; -x_9 = lean_apply_7(x_1, x_3, x_4, x_2, x_5, x_6, x_7, x_8); -return x_9; -} +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_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__13(lean_object* x_1, 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; -x_8 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__12), 8, 2); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_3); -x_9 = l_Lean_Compiler_visitLetImp(x_2, x_8, x_4, x_5, x_6, x_7); -return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_10; -x_10 = l_Lean_Expr_consumeMData(x_4); -switch (lean_obj_tag(x_10)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_12 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_11, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -case 2: -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_13 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_14 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_13, x_5, x_6, x_7, x_8, x_9); -return x_14; -} -case 5: -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_10, 0); -lean_inc(x_15); -switch (lean_obj_tag(x_15)) { -case 0: -{ -lean_object* x_16; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_6 = lean_box(0); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_mkHashMapImp___rarg(x_7); +x_9 = lean_st_ref_get(x_4, x_5); +x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); -x_16 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); +lean_dec(x_9); +x_11 = lean_st_mk_ref(x_8, 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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1; +x_15 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2; +x_16 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3; +lean_inc(x_4); +lean_inc(x_12); +x_17 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2(x_14, x_15, x_16, x_1, x_6, x_12, x_2, x_3, x_4, x_13); if (lean_obj_tag(x_17) == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_3); -lean_dec(x_1); -x_18 = lean_ctor_get(x_16, 1); +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_17, 1); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_18); -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; -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); lean_dec(x_17); -x_22 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_22, 0, x_10); -lean_closure_set(x_22, 1, x_21); -x_23 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_23, 0, x_22); -x_24 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__1), 9, 3); -lean_closure_set(x_24, 0, x_2); -lean_closure_set(x_24, 1, x_1); -lean_closure_set(x_24, 2, x_3); -x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -x_26 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_25, x_5, x_6, x_7, x_8, x_20); -return x_26; -} -} -else -{ -uint8_t x_27; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) +x_19 = lean_st_ref_get(x_4, x_18); +lean_dec(x_4); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_st_ref_get(x_12, x_20); +lean_dec(x_12); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -return x_16; +lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(x_23); +x_25 = lean_array_get_size(x_24); +x_26 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_27 = 0; +x_28 = l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_26, x_27, x_24); +lean_ctor_set(x_21, 0, x_28); +return x_21; } 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_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; lean_object* x_35; lean_object* x_36; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); 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_dec(x_21); +x_31 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(x_29); +x_32 = lean_array_get_size(x_31); +x_33 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_34 = 0; +x_35 = l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_33, x_34, x_31); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_30); +return x_36; } -case 1: -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_3); -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_15, 0); -lean_inc(x_32); -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_33 = lean_apply_7(x_1, x_32, x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_apply_6(x_2, x_31, x_5, x_6, x_7, x_8, x_34); -return x_35; } else { -uint8_t x_36; -lean_dec(x_31); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_36 = !lean_is_exclusive(x_33); -if (x_36 == 0) +uint8_t x_37; +lean_dec(x_12); +lean_dec(x_4); +x_37 = !lean_is_exclusive(x_17); +if (x_37 == 0) { -return x_33; +return x_17; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 0); -x_38 = lean_ctor_get(x_33, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_17, 0); +x_39 = lean_ctor_get(x_17, 1); +lean_inc(x_39); lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_33); -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; +lean_dec(x_17); +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; } } } -case 2: -{ -lean_object* x_40; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_40 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) +} +LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1___boxed(lean_object* x_1) { +_start: { -lean_object* x_42; lean_object* x_43; -lean_dec(x_3); +lean_object* x_2; +x_2 = l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1(x_1); lean_dec(x_1); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_42); -return x_43; +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -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_44 = lean_ctor_get(x_40, 1); -lean_inc(x_44); -lean_dec(x_40); -x_45 = lean_ctor_get(x_41, 0); -lean_inc(x_45); -lean_dec(x_41); -x_46 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_46, 0, x_10); -lean_closure_set(x_46, 1, x_45); -x_47 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_47, 0, x_46); -x_48 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__2), 9, 3); -lean_closure_set(x_48, 0, x_2); -lean_closure_set(x_48, 1, x_1); -lean_closure_set(x_48, 2, x_3); -x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_49, 0, x_47); -lean_closure_set(x_49, 1, x_48); -x_50 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_49, x_5, x_6, x_7, x_8, x_44); -return x_50; +lean_object* x_3; +x_3 = l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(x_1, x_2); +lean_dec(x_2); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_51; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_40); -if (x_51 == 0) +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -return x_40; +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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_4, x_5, x_3); +return x_6; } -else +} +static lean_object* _init_l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1() { +_start: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_40, 0); -x_53 = lean_ctor_get(x_40, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_40); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; +} } +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; lean_object* x_7; lean_object* x_8; +x_6 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1; +x_7 = lean_panic_fn(x_6, x_1); +x_8 = lean_apply_4(x_7, x_2, x_3, x_4, x_5); +return x_8; } } -case 3: +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1() { +_start: { -lean_object* x_55; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_55 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; lean_object* x_58; -lean_dec(x_3); -lean_dec(x_1); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_57); -return x_58; +lean_object* x_1; +x_1 = l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; } -else +} +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2() { +_start: { -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_59 = lean_ctor_get(x_55, 1); -lean_inc(x_59); -lean_dec(x_55); -x_60 = lean_ctor_get(x_56, 0); -lean_inc(x_60); -lean_dec(x_56); -x_61 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_61, 0, x_10); -lean_closure_set(x_61, 1, x_60); -x_62 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_62, 0, x_61); -x_63 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__3), 9, 3); -lean_closure_set(x_63, 0, x_2); -lean_closure_set(x_63, 1, x_1); -lean_closure_set(x_63, 2, x_3); -x_64 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_64, 0, x_62); -lean_closure_set(x_64, 1, x_63); -x_65 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_64, x_5, x_6, x_7, x_8, x_59); -return x_65; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -else -{ -uint8_t x_66; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_66 = !lean_is_exclusive(x_55); -if (x_66 == 0) +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3() { +_start: { -return x_55; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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 -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; -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_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_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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; } } -case 4: -{ -lean_object* x_70; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_70 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_70) == 0) -{ -lean_object* x_71; -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -if (lean_obj_tag(x_71) == 0) +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5() { +_start: { -lean_object* x_72; lean_object* x_73; -lean_dec(x_3); -lean_dec(x_1); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); -x_73 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_72); -return x_73; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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 +} +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6() { +_start: { -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_74 = lean_ctor_get(x_70, 1); -lean_inc(x_74); -lean_dec(x_70); -x_75 = lean_ctor_get(x_71, 0); -lean_inc(x_75); -lean_dec(x_71); -x_76 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_76, 0, x_10); -lean_closure_set(x_76, 1, x_75); -x_77 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_77, 0, x_76); -x_78 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__4), 9, 3); -lean_closure_set(x_78, 0, x_2); -lean_closure_set(x_78, 1, x_1); -lean_closure_set(x_78, 2, x_3); -x_79 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_79, 0, x_77); -lean_closure_set(x_79, 1, x_78); -x_80 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_79, x_5, x_6, x_7, x_8, x_74); -return x_80; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3; +x_3 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4; +x_4 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5; +x_5 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set(x_5, 4, x_4); +lean_ctor_set(x_5, 5, x_2); +lean_ctor_set(x_5, 6, x_3); +lean_ctor_set(x_5, 7, x_3); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_81; -lean_dec(x_10); -lean_dec(x_8); +lean_object* x_6; lean_object* x_7; 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_6 = lean_ctor_get(x_3, 5); +x_7 = lean_st_ref_get(x_4, x_5); +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); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_81 = !lean_is_exclusive(x_70); -if (x_81 == 0) +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_st_ref_get(x_4, x_9); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -return x_70; +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_13, 0); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_3, 2); +x_18 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6; +lean_inc(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_18); +lean_ctor_set(x_19, 2, x_16); +lean_ctor_set(x_19, 3, x_17); +x_20 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_1); +lean_inc(x_6); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_6); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_21); +return x_13; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_70, 0); -x_83 = lean_ctor_get(x_70, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_70); -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_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_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_3, 2); +x_26 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6; +lean_inc(x_25); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_25); +x_28 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_1); +lean_inc(x_6); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_23); +return x_30; } } } -case 5: +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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: { -lean_object* x_85; -lean_dec(x_15); +if (lean_obj_tag(x_2) == 8) +{ +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; +x_8 = lean_ctor_get(x_2, 0); lean_inc(x_8); -lean_inc(x_7); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_2, 2); lean_inc(x_10); -x_85 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_85) == 0) +x_11 = lean_ctor_get(x_2, 3); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*4 + 8); +lean_dec(x_2); +x_13 = lean_expr_instantiate_rev(x_9, x_3); +lean_dec(x_9); +x_14 = lean_expr_instantiate_rev(x_10, x_3); +lean_dec(x_10); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +lean_inc(x_14); +x_15 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_14, x_1, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_86; -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -if (lean_obj_tag(x_86) == 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_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_Compiler_mkLetDecl(x_8, x_13, x_14, x_12, x_4, x_5, 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_array_push(x_3, x_18); +x_2 = x_11; +x_3 = x_20; +x_7 = x_19; +goto _start; +} +else { -lean_object* x_87; lean_object* x_88; +uint8_t x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(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_1); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -lean_dec(x_85); -x_88 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_87); -return x_88; +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +return x_15; } 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; -x_89 = lean_ctor_get(x_85, 1); -lean_inc(x_89); -lean_dec(x_85); -x_90 = lean_ctor_get(x_86, 0); -lean_inc(x_90); -lean_dec(x_86); -x_91 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_91, 0, x_10); -lean_closure_set(x_91, 1, x_90); -x_92 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_92, 0, x_91); -x_93 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__5), 9, 3); -lean_closure_set(x_93, 0, x_2); -lean_closure_set(x_93, 1, x_1); -lean_closure_set(x_93, 2, x_3); -x_94 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_94, 0, x_92); -lean_closure_set(x_94, 1, x_93); -x_95 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_94, x_5, x_6, x_7, x_8, x_89); -return x_95; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +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 { -uint8_t x_96; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_26; lean_object* x_27; lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_26 = lean_expr_instantiate_rev(x_2, x_3); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_96 = !lean_is_exclusive(x_85); -if (x_96 == 0) -{ -return x_85; -} -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_85, 0); -x_98 = lean_ctor_get(x_85, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_85); -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; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_7); +return x_27; } } } -case 6: +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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) { +_start: { -lean_object* x_100; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_100 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_100) == 0) +lean_object* x_7; +x_7 = l_Lean_Expr_consumeMData(x_1); +lean_dec(x_1); +switch (lean_obj_tag(x_7)) { +case 0: { -lean_object* x_101; -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -if (lean_obj_tag(x_101) == 0) +lean_object* x_8; lean_object* x_9; +lean_dec(x_7); +lean_dec(x_2); +x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_9 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_8, x_3, x_4, x_5, x_6); +return x_9; +} +case 1: { -lean_object* x_102; lean_object* x_103; -lean_dec(x_3); -lean_dec(x_1); -x_102 = lean_ctor_get(x_100, 1); -lean_inc(x_102); -lean_dec(x_100); -x_103 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_102); -return x_103; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_7, 0); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_apply_5(x_2, x_10, x_3, x_4, x_5, x_6); +return x_11; } -else +case 2: { -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_100, 1); -lean_inc(x_104); -lean_dec(x_100); -x_105 = lean_ctor_get(x_101, 0); -lean_inc(x_105); -lean_dec(x_101); -x_106 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_106, 0, x_10); -lean_closure_set(x_106, 1, x_105); -x_107 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_107, 0, x_106); -x_108 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__6), 9, 3); -lean_closure_set(x_108, 0, x_2); -lean_closure_set(x_108, 1, x_1); -lean_closure_set(x_108, 2, x_3); -x_109 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_109, 0, x_107); -lean_closure_set(x_109, 1, x_108); -x_110 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_109, x_5, x_6, x_7, x_8, x_104); -return x_110; +lean_object* x_12; lean_object* x_13; +lean_dec(x_7); +lean_dec(x_2); +x_12 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_13 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_12, x_3, x_4, x_5, x_6); +return x_13; } +case 5: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_7, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); +lean_dec(x_7); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_16 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_14, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_1 = x_15; +x_6 = x_17; +goto _start; } else { -uint8_t x_111; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_19; +lean_dec(x_15); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_111 = !lean_is_exclusive(x_100); -if (x_111 == 0) +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -return x_100; +return x_16; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_112 = lean_ctor_get(x_100, 0); -x_113 = lean_ctor_get(x_100, 1); -lean_inc(x_113); -lean_inc(x_112); -lean_dec(x_100); -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_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_16, 0); +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_16); +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; } } } -case 7: -{ -lean_object* x_115; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_115 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_115) == 0) +case 6: { -lean_object* x_116; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -if (lean_obj_tag(x_116) == 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; lean_object* x_43; +x_23 = lean_st_ref_get(x_5, x_6); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_get(x_3, x_24); +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_st_ref_get(x_5, x_27); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_st_ref_take(x_3, 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); +lean_dec(x_30); +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_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_34); +x_37 = lean_st_ref_set(x_3, x_36, x_32); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Compiler_visitLambda(x_7, x_3, x_4, x_5, x_38); +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_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +lean_inc(x_5); +lean_inc(x_3); +x_43 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_42, x_2, x_3, x_4, x_5, x_41); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_117; lean_object* x_118; +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; uint8_t x_58; +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_st_ref_get(x_5, x_45); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_st_ref_get(x_3, x_47); +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 = lean_ctor_get(x_26, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_26, 1); +lean_inc(x_52); +lean_dec(x_26); +x_53 = lean_ctor_get(x_49, 2); +lean_inc(x_53); +lean_dec(x_49); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +lean_ctor_set(x_54, 2, x_53); +x_55 = lean_st_ref_get(x_5, x_50); +lean_dec(x_5); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = lean_st_ref_set(x_3, x_54, x_56); lean_dec(x_3); -lean_dec(x_1); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -lean_dec(x_115); -x_118 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_117); -return x_118; +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) +{ +lean_object* x_59; +x_59 = lean_ctor_get(x_57, 0); +lean_dec(x_59); +lean_ctor_set(x_57, 0, x_44); +return x_57; } 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_ctor_get(x_115, 1); -lean_inc(x_119); -lean_dec(x_115); -x_120 = lean_ctor_get(x_116, 0); -lean_inc(x_120); -lean_dec(x_116); -x_121 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_121, 0, x_10); -lean_closure_set(x_121, 1, x_120); -x_122 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_122, 0, x_121); -x_123 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__7), 9, 3); -lean_closure_set(x_123, 0, x_2); -lean_closure_set(x_123, 1, x_1); -lean_closure_set(x_123, 2, x_3); -x_124 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_124, 0, x_122); -lean_closure_set(x_124, 1, x_123); -x_125 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_124, x_5, x_6, x_7, x_8, x_119); -return x_125; +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +lean_dec(x_57); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_44); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } else { -uint8_t x_126; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +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; uint8_t x_76; +x_62 = lean_ctor_get(x_43, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_43, 1); +lean_inc(x_63); +lean_dec(x_43); +x_64 = lean_st_ref_get(x_5, x_63); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = lean_st_ref_get(x_3, x_65); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_26, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_26, 1); +lean_inc(x_70); +lean_dec(x_26); +x_71 = lean_ctor_get(x_67, 2); +lean_inc(x_71); +lean_dec(x_67); +x_72 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_70); +lean_ctor_set(x_72, 2, x_71); +x_73 = lean_st_ref_get(x_5, x_68); lean_dec(x_5); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = lean_st_ref_set(x_3, x_72, x_74); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_126 = !lean_is_exclusive(x_115); -if (x_126 == 0) +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) { -return x_115; +lean_object* x_77; +x_77 = lean_ctor_get(x_75, 0); +lean_dec(x_77); +lean_ctor_set_tag(x_75, 1); +lean_ctor_set(x_75, 0, x_62); +return x_75; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_115, 0); -x_128 = lean_ctor_get(x_115, 1); -lean_inc(x_128); -lean_inc(x_127); -lean_dec(x_115); -x_129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -return x_129; +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_dec(x_75); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_62); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } case 8: { -lean_object* x_130; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_130 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_130) == 0) +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; +x_80 = lean_st_ref_get(x_5, x_6); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_st_ref_get(x_3, x_81); +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_st_ref_get(x_5, x_84); +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +x_87 = lean_st_ref_take(x_3, x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_ctor_get(x_88, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_88, 2); +lean_inc(x_91); +lean_dec(x_88); +x_92 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_93 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +x_94 = lean_st_ref_set(x_3, x_93, x_89); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_96 = l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(x_2, x_7, x_92, x_3, x_4, x_5, x_95); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_131; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -if (lean_obj_tag(x_131) == 0) +lean_object* x_97; lean_object* x_98; lean_object* x_99; +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); +lean_inc(x_5); +lean_inc(x_3); +x_99 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_97, x_2, x_3, x_4, x_5, x_98); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_132; lean_object* x_133; +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; uint8_t x_114; +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_102 = lean_st_ref_get(x_5, x_101); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_104 = lean_st_ref_get(x_3, x_103); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_ctor_get(x_83, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_83, 1); +lean_inc(x_108); +lean_dec(x_83); +x_109 = lean_ctor_get(x_105, 2); +lean_inc(x_109); +lean_dec(x_105); +x_110 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +lean_ctor_set(x_110, 2, x_109); +x_111 = lean_st_ref_get(x_5, x_106); +lean_dec(x_5); +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +lean_dec(x_111); +x_113 = lean_st_ref_set(x_3, x_110, x_112); lean_dec(x_3); -lean_dec(x_1); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -lean_dec(x_130); -x_133 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_132); -return x_133; +x_114 = !lean_is_exclusive(x_113); +if (x_114 == 0) +{ +lean_object* x_115; +x_115 = lean_ctor_get(x_113, 0); +lean_dec(x_115); +lean_ctor_set(x_113, 0, x_100); +return x_113; } 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; -x_134 = lean_ctor_get(x_130, 1); -lean_inc(x_134); -lean_dec(x_130); -x_135 = lean_ctor_get(x_131, 0); -lean_inc(x_135); -lean_dec(x_131); -x_136 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_136, 0, x_10); -lean_closure_set(x_136, 1, x_135); -x_137 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_137, 0, x_136); -x_138 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__8), 9, 3); -lean_closure_set(x_138, 0, x_2); -lean_closure_set(x_138, 1, x_1); -lean_closure_set(x_138, 2, x_3); -x_139 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_139, 0, x_137); -lean_closure_set(x_139, 1, x_138); -x_140 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_139, x_5, x_6, x_7, x_8, x_134); -return x_140; +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_113, 1); +lean_inc(x_116); +lean_dec(x_113); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_100); +lean_ctor_set(x_117, 1, x_116); +return x_117; } } else { -uint8_t x_141; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_141 = !lean_is_exclusive(x_130); -if (x_141 == 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; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_118 = lean_ctor_get(x_99, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_99, 1); +lean_inc(x_119); +lean_dec(x_99); +x_120 = lean_st_ref_get(x_5, x_119); +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +x_122 = lean_st_ref_get(x_3, x_121); +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_83, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_83, 1); +lean_inc(x_126); +lean_dec(x_83); +x_127 = lean_ctor_get(x_123, 2); +lean_inc(x_127); +lean_dec(x_123); +x_128 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +lean_ctor_set(x_128, 2, x_127); +x_129 = lean_st_ref_get(x_5, x_124); +lean_dec(x_5); +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +lean_dec(x_129); +x_131 = lean_st_ref_set(x_3, x_128, x_130); +lean_dec(x_3); +x_132 = !lean_is_exclusive(x_131); +if (x_132 == 0) { -return x_130; +lean_object* x_133; +x_133 = lean_ctor_get(x_131, 0); +lean_dec(x_133); +lean_ctor_set_tag(x_131, 1); +lean_ctor_set(x_131, 0, x_118); +return x_131; } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_130, 0); -x_143 = lean_ctor_get(x_130, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_130); -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; +lean_object* x_134; lean_object* x_135; +x_134 = lean_ctor_get(x_131, 1); +lean_inc(x_134); +lean_dec(x_131); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_118); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } -case 9: -{ -lean_object* x_145; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_145 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_145) == 0) -{ -lean_object* x_146; -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -if (lean_obj_tag(x_146) == 0) +else { -lean_object* x_147; lean_object* x_148; +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; uint8_t x_150; +lean_dec(x_4); +lean_dec(x_2); +x_136 = lean_ctor_get(x_96, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_96, 1); +lean_inc(x_137); +lean_dec(x_96); +x_138 = lean_st_ref_get(x_5, x_137); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +lean_dec(x_138); +x_140 = lean_st_ref_get(x_3, x_139); +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_140, 1); +lean_inc(x_142); +lean_dec(x_140); +x_143 = lean_ctor_get(x_83, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_83, 1); +lean_inc(x_144); +lean_dec(x_83); +x_145 = lean_ctor_get(x_141, 2); +lean_inc(x_145); +lean_dec(x_141); +x_146 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_146, 0, x_143); +lean_ctor_set(x_146, 1, x_144); +lean_ctor_set(x_146, 2, x_145); +x_147 = lean_st_ref_get(x_5, x_142); +lean_dec(x_5); +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +lean_dec(x_147); +x_149 = lean_st_ref_set(x_3, x_146, x_148); lean_dec(x_3); -lean_dec(x_1); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_147); -return x_148; +x_150 = !lean_is_exclusive(x_149); +if (x_150 == 0) +{ +lean_object* x_151; +x_151 = lean_ctor_get(x_149, 0); +lean_dec(x_151); +lean_ctor_set_tag(x_149, 1); +lean_ctor_set(x_149, 0, x_136); +return x_149; } 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; lean_object* x_155; -x_149 = lean_ctor_get(x_145, 1); -lean_inc(x_149); -lean_dec(x_145); -x_150 = lean_ctor_get(x_146, 0); -lean_inc(x_150); -lean_dec(x_146); -x_151 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_151, 0, x_10); -lean_closure_set(x_151, 1, x_150); -x_152 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_152, 0, x_151); -x_153 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__9), 9, 3); -lean_closure_set(x_153, 0, x_2); -lean_closure_set(x_153, 1, x_1); -lean_closure_set(x_153, 2, x_3); -x_154 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_154, 0, x_152); -lean_closure_set(x_154, 1, x_153); -x_155 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_154, x_5, x_6, x_7, x_8, x_149); -return x_155; +lean_object* x_152; lean_object* x_153; +x_152 = lean_ctor_get(x_149, 1); +lean_inc(x_152); +lean_dec(x_149); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_136); +lean_ctor_set(x_153, 1, x_152); +return x_153; } } -else +} +case 10: { -uint8_t x_156; -lean_dec(x_10); -lean_dec(x_8); +lean_object* x_154; lean_object* x_155; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_156 = !lean_is_exclusive(x_145); -if (x_156 == 0) +x_154 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; +x_155 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_154, x_3, x_4, x_5, x_6); +return x_155; +} +case 11: { -return x_145; +lean_object* x_156; +x_156 = lean_ctor_get(x_7, 2); +lean_inc(x_156); +lean_dec(x_7); +x_1 = x_156; +goto _start; } -else +default: { -lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_157 = lean_ctor_get(x_145, 0); -x_158 = lean_ctor_get(x_145, 1); -lean_inc(x_158); -lean_inc(x_157); -lean_dec(x_145); -x_159 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); +lean_object* x_158; lean_object* x_159; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_158 = lean_box(0); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_6); return x_159; } } } -case 10: +} +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1() { +_start: { -lean_object* x_160; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_160 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_160) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointChecker.containsNoJp", 54); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2() { +_start: { -lean_object* x_161; -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -if (lean_obj_tag(x_161) == 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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; +x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1; +x_3 = lean_unsigned_to_nat(157u); +x_4 = lean_unsigned_to_nat(39u); +x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3() { +_start: { -lean_object* x_162; lean_object* x_163; -lean_dec(x_3); -lean_dec(x_1); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); -x_163 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_162); -return x_163; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Join point ", 11); +return x_1; } -else +} +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4() { +_start: { -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_164 = lean_ctor_get(x_160, 1); -lean_inc(x_164); -lean_dec(x_160); -x_165 = lean_ctor_get(x_161, 0); -lean_inc(x_165); -lean_dec(x_161); -x_166 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_166, 0, x_10); -lean_closure_set(x_166, 1, x_165); -x_167 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_167, 0, x_166); -x_168 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__10), 9, 3); -lean_closure_set(x_168, 0, x_2); -lean_closure_set(x_168, 1, x_1); -lean_closure_set(x_168, 2, x_3); -x_169 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_169, 0, x_167); -lean_closure_set(x_169, 1, x_168); -x_170 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_169, x_5, x_6, x_7, x_8, x_164); -return x_170; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" in forbidden position", 22); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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_171; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_171 = !lean_is_exclusive(x_160); -if (x_171 == 0) +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(x_1, x_2, x_3, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +if (lean_obj_tag(x_7) == 0) { -return x_160; +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2; +x_10 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_9, x_2, x_3, x_4, x_8); +return x_10; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_172 = lean_ctor_get(x_160, 0); -x_173 = lean_ctor_get(x_160, 1); -lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_160); -x_174 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_173); -return x_174; -} -} -} -default: +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) { -lean_object* x_175; -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_10); -x_175 = l_Lean_Compiler_isCasesApp_x3f(x_10, x_7, x_8, x_9); -if (lean_obj_tag(x_175) == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_6, 1); +x_13 = lean_ctor_get(x_6, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_7, 0); +lean_inc(x_14); +lean_dec(x_7); +x_15 = l_Lean_LocalDecl_isJp(x_14); +if (x_15 == 0) { -lean_object* x_176; -x_176 = lean_ctor_get(x_175, 0); -lean_inc(x_176); -if (lean_obj_tag(x_176) == 0) +lean_object* x_16; +lean_dec(x_14); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_box(0); +lean_ctor_set(x_6, 0, x_16); +return x_6; +} +else { -lean_object* x_177; lean_object* x_178; +lean_object* x_17; uint8_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; +lean_free_object(x_6); +x_17 = l_Lean_LocalDecl_userName(x_14); +lean_dec(x_14); +x_18 = 1; +x_19 = l_Lean_Name_toString(x_17, x_18); +x_20 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; +x_21 = lean_string_append(x_20, x_19); +lean_dec(x_19); +x_22 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4; +x_23 = lean_string_append(x_21, x_22); +x_24 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_25, x_2, x_3, x_4, x_12); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -x_177 = lean_ctor_get(x_175, 1); -lean_inc(x_177); -lean_dec(x_175); -x_178 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_177); -return x_178; +lean_dec(x_2); +return x_26; +} } 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; -x_179 = lean_ctor_get(x_175, 1); -lean_inc(x_179); -lean_dec(x_175); -x_180 = lean_ctor_get(x_176, 0); -lean_inc(x_180); -lean_dec(x_176); -x_181 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_181, 0, x_10); -lean_closure_set(x_181, 1, x_180); -x_182 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_182, 0, x_181); -x_183 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__11), 9, 3); -lean_closure_set(x_183, 0, x_2); -lean_closure_set(x_183, 1, x_1); -lean_closure_set(x_183, 2, x_3); -x_184 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_184, 0, x_182); -lean_closure_set(x_184, 1, x_183); -x_185 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_184, x_5, x_6, x_7, x_8, x_179); -return x_185; -} -} -else -{ -uint8_t x_186; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_6, 1); +lean_inc(x_27); lean_dec(x_6); -lean_dec(x_5); +x_28 = lean_ctor_get(x_7, 0); +lean_inc(x_28); +lean_dec(x_7); +x_29 = l_Lean_LocalDecl_isJp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_28); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_186 = !lean_is_exclusive(x_175); -if (x_186 == 0) -{ -return x_175; +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_27); +return x_31; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_175, 0); -x_188 = lean_ctor_get(x_175, 1); -lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_175); -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; +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; +x_32 = l_Lean_LocalDecl_userName(x_28); +lean_dec(x_28); +x_33 = 1; +x_34 = l_Lean_Name_toString(x_32, x_33); +x_35 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; +x_36 = lean_string_append(x_35, x_34); +lean_dec(x_34); +x_37 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4; +x_38 = lean_string_append(x_36, x_37); +x_39 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_40, x_2, x_3, x_4, x_27); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_41; } } } } } -case 6: +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1() { +_start: { -lean_object* x_190; -lean_dec(x_3); -lean_dec(x_1); -x_190 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_9); -return x_190; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1), 5, 0); +return x_1; } -case 8: -{ -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; -lean_inc(x_3); -x_191 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___lambda__13), 7, 2); -lean_closure_set(x_191, 0, x_3); -lean_closure_set(x_191, 1, x_10); -x_192 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1___boxed), 9, 3); -lean_closure_set(x_192, 0, x_1); -lean_closure_set(x_192, 1, x_2); -lean_closure_set(x_192, 2, x_3); -x_193 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_193, 0, x_191); -lean_closure_set(x_193, 1, x_192); -x_194 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_193, x_5, x_6, x_7, x_8, x_9); -return x_194; } -case 10: +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_195; lean_object* x_196; -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_195 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_196 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2(x_195, x_5, x_6, x_7, x_8, x_9); -return x_196; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1; +x_7 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_1, x_6, x_2, x_3, x_4, x_5); +return x_7; } -case 11: -{ -lean_object* x_197; -lean_dec(x_3); -lean_dec(x_1); -x_197 = lean_apply_6(x_2, x_10, x_5, x_6, x_7, x_8, x_9); -return x_197; } -default: +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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_198; lean_object* x_199; -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_6; +x_6 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_198 = lean_box(0); -x_199 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_199, 0, x_198); -lean_ctor_set(x_199, 1, x_9); -return x_199; -} -} +return x_6; } } -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__49(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1() { _start: { -if (lean_obj_tag(x_2) == 0) -{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointChecker.checkJoinPoints.goTailApp", 67); 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; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; -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_Lean_Name_hash___override(x_4); -x_8 = lean_uint64_to_usize(x_7); -x_9 = lean_usize_modn(x_8, x_6); -lean_dec(x_6); -x_10 = lean_array_uget(x_1, x_9); -lean_ctor_set(x_2, 2, x_10); -x_11 = lean_array_uset(x_1, x_9, x_2); -x_1 = x_11; -x_2 = x_5; -goto _start; } -else +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2() { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* 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_13 = lean_ctor_get(x_2, 0); -x_14 = lean_ctor_get(x_2, 1); -x_15 = lean_ctor_get(x_2, 2); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_2); -x_16 = lean_array_get_size(x_1); -x_17 = l_Lean_Name_hash___override(x_13); -x_18 = lean_uint64_to_usize(x_17); -x_19 = lean_usize_modn(x_18, x_16); -lean_dec(x_16); -x_20 = lean_array_uget(x_1, x_19); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_14); -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_15; -goto _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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; +x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1; +x_3 = lean_unsigned_to_nat(178u); +x_4 = lean_unsigned_to_nat(39u); +x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___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_Std_HashMapImp_moveEntries___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__48(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" : ", 3); +return x_1; } -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_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__49(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; } +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" is not fully applied in ", 25); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__47(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5() { _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_array_get_size(x_2); -x_4 = lean_unsigned_to_nat(2u); -x_5 = lean_nat_mul(x_3, x_4); -lean_dec(x_3); -x_6 = lean_box(0); -x_7 = lean_mk_array(x_5, x_6); -x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Std_HashMapImp_moveEntries___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__48(x_8, x_2, x_7); -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_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp(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_3) == 0) +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(x_1, x_3, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2; +x_11 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_10, x_3, x_4, x_5, x_9); +return x_11; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_7); +if (x_12 == 0) { -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 = lean_name_eq(x_6, x_1); -if (x_9 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_7, 1); +x_14 = lean_ctor_get(x_7, 0); +lean_dec(x_14); +x_15 = lean_ctor_get(x_8, 0); +lean_inc(x_15); +lean_dec(x_8); +x_16 = l_Lean_LocalDecl_isJp(x_15); +if (x_16 == 0) { -lean_object* x_10; -x_10 = l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_10); -return x_3; +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_17 = lean_box(0); +lean_ctor_set(x_7, 0, x_17); +return x_7; } else { -lean_dec(x_7); -lean_dec(x_6); -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; uint8_t x_21; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); +x_20 = l_Lean_Compiler_jpArity(x_15); +x_21 = lean_nat_dec_eq(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (x_21 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_22; uint8_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; 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_free_object(x_7); +x_22 = l_Lean_LocalDecl_userName(x_15); +x_23 = 1; +x_24 = l_Lean_Name_toString(x_22, x_23); +x_25 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; +x_26 = lean_string_append(x_25, x_24); +lean_dec(x_24); +x_27 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3; +x_28 = lean_string_append(x_26, x_27); +x_29 = l_Lean_LocalDecl_type(x_15); +lean_dec(x_15); +x_30 = lean_expr_dbg_to_string(x_29); +lean_dec(x_29); +x_31 = lean_string_append(x_28, x_30); +lean_dec(x_30); +x_32 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4; +x_33 = lean_string_append(x_31, x_32); +x_34 = lean_expr_dbg_to_string(x_2); +x_35 = lean_string_append(x_33, x_34); +lean_dec(x_34); +x_36 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5; +x_37 = lean_string_append(x_35, x_36); +x_38 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_39, x_3, x_4, x_5, x_13); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_14 = lean_name_eq(x_11, x_1); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(x_1, x_2, x_13); -x_16 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_12); -lean_ctor_set(x_16, 2, x_15); -return x_16; +return x_40; } else { -lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 2, x_13); -return x_17; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -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; uint64_t x_8; size_t x_9; size_t x_10; lean_object* x_11; uint8_t x_12; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_array_get_size(x_6); -x_8 = l_Lean_Name_hash___override(x_2); -x_9 = lean_uint64_to_usize(x_8); -x_10 = lean_usize_modn(x_9, x_7); -x_11 = lean_array_uget(x_6, x_10); -x_12 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_2, x_11); -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_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_5, x_13); -lean_dec(x_5); -x_15 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_15, 0, x_2); -lean_ctor_set(x_15, 1, x_3); -lean_ctor_set(x_15, 2, x_11); -x_16 = lean_array_uset(x_6, x_10, x_15); -x_17 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_14); -x_18 = lean_nat_dec_le(x_17, x_7); -lean_dec(x_7); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; -lean_free_object(x_1); -x_19 = l_Std_HashMapImp_expand___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__47(x_14, x_16); -return x_19; -} -else -{ -lean_ctor_set(x_1, 1, x_16); -lean_ctor_set(x_1, 0, x_14); -return x_1; -} -} -else -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_7); -x_20 = l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(x_2, x_3, x_11); -x_21 = lean_array_uset(x_6, x_10, 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; uint64_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; -x_22 = lean_ctor_get(x_1, 0); -x_23 = lean_ctor_get(x_1, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_1); -x_24 = lean_array_get_size(x_23); -x_25 = l_Lean_Name_hash___override(x_2); -x_26 = lean_uint64_to_usize(x_25); -x_27 = lean_usize_modn(x_26, x_24); -x_28 = lean_array_uget(x_23, x_27); -x_29 = l_Std_AssocList_contains___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4(x_2, x_28); -if (x_29 == 0) -{ -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_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_22, x_30); -lean_dec(x_22); -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_2); -lean_ctor_set(x_32, 1, x_3); -lean_ctor_set(x_32, 2, x_28); -x_33 = lean_array_uset(x_23, x_27, x_32); -x_34 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_31); -x_35 = lean_nat_dec_le(x_34, x_24); -lean_dec(x_24); -lean_dec(x_34); -if (x_35 == 0) -{ -lean_object* x_36; -x_36 = l_Std_HashMapImp_expand___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__47(x_31, x_33); -return x_36; -} -else -{ -lean_object* x_37; -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_31); -lean_ctor_set(x_37, 1, x_33); -return x_37; -} -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_24); -x_38 = l_Std_AssocList_replace___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__50(x_2, x_3, x_28); -x_39 = lean_array_uset(x_23, x_27, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_22); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___boxed), 7, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn), 6, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue), 7, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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; 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_2, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_2, 1); -lean_inc(x_9); -lean_dec(x_2); -x_10 = lean_st_ref_get(x_6, x_7); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_st_ref_take(x_3, 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_array_get_size(x_8); -lean_dec(x_8); -x_16 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_13, x_1, x_15); -x_17 = lean_st_ref_set(x_3, x_16, x_14); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1; -x_20 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2; -x_21 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3; -x_22 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_19, x_20, x_21, x_9, x_3, x_4, x_5, x_6, x_18); -lean_dec(x_9); -return x_22; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue(lean_object* x_1, 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_2) == 6) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -lean_inc(x_2); -x_8 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_8, 0, x_2); -x_9 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2), 7, 1); -lean_closure_set(x_10, 0, x_1); -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_11, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_2); -return x_12; -} -else -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -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; -} -} -else -{ -uint8_t x_17; -lean_dec(x_2); -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_object* x_24; -lean_dec(x_1); -x_21 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1; -x_22 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2; -x_23 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3; -x_24 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_21, x_22, x_23, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -lean_ctor_set(x_24, 0, x_2); -return x_24; -} -else -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_2); -lean_ctor_set(x_28, 1, x_27); -return x_28; -} -} -else -{ -uint8_t x_29; -lean_dec(x_2); -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) -{ -return x_24; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_24, 0); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_24); -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_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__3(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_10) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__4(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__5(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__6(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__7(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__8(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__9(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__10(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__11(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__12(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__13(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__14(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__15(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__16(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__17(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__18(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__19(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_10) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__20(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__21(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__22(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__23(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__24(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__25(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__26(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__27(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__28(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__29(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__30(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__31(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__32(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__33(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__34(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__35(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__36(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__37(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__38(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__39(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__40(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__41(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__42(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__43(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__44(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); -lean_dec(x_5); -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__45(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_15; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___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_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Std_mkHashMapImp___rarg(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_4, 1); -x_11 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__1(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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; lean_object* x_14; -x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_10, 0, x_4); -x_11 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_11, 0, x_10); -x_12 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2___lambda__1___boxed), 9, 3); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_3); -x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__7___rarg), 7, 2); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_12); -x_14 = l_Lean_Compiler_withNewScope___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__8(x_13, x_5, x_6, x_7, x_8, x_9); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(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_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(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; -} -} -} -static lean_object* _init_l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___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_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(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); -lean_inc(x_2); -lean_dec(x_1); -x_3 = lean_array_get_size(x_2); -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_lt(x_4, x_3); -if (x_5 == 0) -{ -lean_object* x_6; -lean_dec(x_3); -lean_dec(x_2); -x_6 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -return x_6; -} -else -{ -uint8_t x_7; -x_7 = lean_nat_dec_le(x_3, x_3); -if (x_7 == 0) -{ -lean_object* x_8; -lean_dec(x_3); -lean_dec(x_2); -x_8 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -return x_8; -} -else -{ -size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; -x_9 = 0; -x_10 = lean_usize_of_nat(x_3); -lean_dec(x_3); -x_11 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(x_2, x_9, x_10, x_11); -lean_dec(x_2); -return x_12; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints(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; -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Std_mkHashMapImp___rarg(x_6); -x_8 = lean_st_ref_get(x_4, x_5); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -lean_dec(x_8); -x_10 = lean_st_mk_ref(x_7, x_9); -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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1; -x_14 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2; -x_15 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3; -lean_inc(x_4); -lean_inc(x_11); -x_16 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2(x_13, x_14, x_15, x_1, x_11, x_2, x_3, x_4, x_12); -if (lean_obj_tag(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_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_st_ref_get(x_4, x_17); -lean_dec(x_4); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_st_ref_get(x_11, x_19); -lean_dec(x_11); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_20, 0); -x_23 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(x_22); -x_24 = lean_array_get_size(x_23); -x_25 = lean_usize_of_nat(x_24); -lean_dec(x_24); -x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_25, x_26, x_23); -lean_ctor_set(x_20, 0, x_27); -return x_20; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_ctor_get(x_20, 0); -x_29 = lean_ctor_get(x_20, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_20); -x_30 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3(x_28); -x_31 = lean_array_get_size(x_30); -x_32 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_33 = 0; -x_34 = l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_32, x_33, x_30); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_29); -return x_35; -} -} -else -{ -uint8_t x_36; -lean_dec(x_11); -lean_dec(x_4); -x_36 = !lean_is_exclusive(x_16); -if (x_36 == 0) -{ -return x_16; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_16, 0); -x_38 = lean_ctor_get(x_16, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_16); -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; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Std_mkHashMap___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__1(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2___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___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_AssocList_foldlM___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__4(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5___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_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__5(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___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; -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_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__6(x_4, x_5, x_3); -return x_6; -} -} -static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_inheritedTraceOptions; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_st_ref_get(x_4, x_5); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1; -x_9 = lean_st_ref_get(x_8, 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; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_3, 2); -x_13 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_11, x_12, x_1); -lean_dec(x_11); -x_14 = lean_box(x_13); -lean_ctor_set(x_9, 0, x_14); -return x_9; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t 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_ctor_get(x_3, 2); -x_18 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_15, x_17, x_1); -lean_dec(x_15); -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_16); -return x_20; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; uint8_t x_9; -x_6 = lean_st_ref_get(x_4, x_5); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_st_ref_get(x_2, x_7); -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, 0); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_local_ctx_find(x_11, x_1); -lean_ctor_set(x_8, 0, 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 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_local_ctx_find(x_15, x_1); -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; -} -} -} -static lean_object* _init_l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabited___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__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; -x_6 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1; -x_7 = lean_panic_fn(x_6, x_1); -x_8 = lean_apply_4(x_7, x_2, x_3, x_4, x_5); -return x_8; -} -} -static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; -} -} -static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3; -x_3 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4; -x_4 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5; -x_5 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_1); -lean_ctor_set(x_5, 2, x_2); -lean_ctor_set(x_5, 3, x_3); -lean_ctor_set(x_5, 4, x_4); -lean_ctor_set(x_5, 5, x_2); -lean_ctor_set(x_5, 6, x_3); -lean_ctor_set(x_5, 7, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; lean_object* x_7; 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_6 = lean_ctor_get(x_3, 5); -x_7 = lean_st_ref_get(x_4, x_5); -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_8, 0); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_st_ref_get(x_4, x_9); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_2, x_12); -x_14 = !lean_is_exclusive(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; lean_object* x_21; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_3, 2); -x_18 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6; -lean_inc(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_18); -lean_ctor_set(x_19, 2, x_16); -lean_ctor_set(x_19, 3, x_17); -x_20 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_1); -lean_inc(x_6); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_6); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set_tag(x_13, 1); -lean_ctor_set(x_13, 0, x_21); -return x_13; -} -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; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_3, 2); -x_26 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6; -lean_inc(x_25); -x_27 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_27, 0, x_10); -lean_ctor_set(x_27, 1, x_26); -lean_ctor_set(x_27, 2, x_24); -lean_ctor_set(x_27, 3, x_25); -x_28 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_1); -lean_inc(x_6); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_6); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_23); -return x_30; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__6(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_Compiler_withNewScopeImp___rarg(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___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_ctor_get(x_2, 1); -lean_inc(x_7); -lean_dec(x_2); -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_7, x_1, x_3, x_4, x_5, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___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_inc(x_3); -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_3, x_1, 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; -x_10 = lean_ctor_get(x_8, 0); -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_3); -return x_8; -} -else -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -} -else -{ -uint8_t x_13; -lean_dec(x_3); -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_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___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; -x_7 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_2, x_1, x_3, x_4, x_5, x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Expr_consumeMData(x_1); -lean_dec(x_1); -switch (lean_obj_tag(x_7)) { -case 0: -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_7); -lean_dec(x_2); -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_9 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_8, x_3, x_4, x_5, x_6); -return x_9; -} -case 1: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_7, 0); -lean_inc(x_10); -lean_dec(x_7); -x_11 = lean_apply_5(x_2, x_10, x_3, x_4, x_5, x_6); -return x_11; -} -case 2: -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_7); -lean_dec(x_2); -x_12 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_13 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_12, x_3, x_4, x_5, x_6); -return x_13; -} -case 5: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_7, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_7, 1); -lean_inc(x_15); -lean_dec(x_7); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_16 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_14, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_1 = x_15; -x_6 = x_17; -goto _start; -} -else -{ -uint8_t x_19; -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -return x_16; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -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; -} -} -} -case 6: -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_23, 0, x_7); -x_24 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__1), 6, 1); -lean_closure_set(x_24, 0, x_2); -x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -x_26 = l_Lean_Compiler_withNewScopeImp___rarg(x_25, x_3, x_4, x_5, x_6); -return x_26; -} -case 8: -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_inc(x_2); -x_27 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__2___boxed), 7, 1); -lean_closure_set(x_27, 0, x_2); -x_28 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLetImp), 6, 2); -lean_closure_set(x_28, 0, x_7); -lean_closure_set(x_28, 1, x_27); -x_29 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__3), 6, 1); -lean_closure_set(x_29, 0, x_2); -x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_30, 0, x_28); -lean_closure_set(x_30, 1, x_29); -x_31 = l_Lean_Compiler_withNewScopeImp___rarg(x_30, x_3, x_4, x_5, x_6); -return x_31; -} -case 10: -{ -lean_object* x_32; lean_object* x_33; -lean_dec(x_7); -lean_dec(x_2); -x_32 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__4; -x_33 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_32, x_3, x_4, x_5, x_6); -return x_33; -} -case 11: -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_7, 2); -lean_inc(x_34); -lean_dec(x_7); -x_1 = x_34; -goto _start; -} -default: -{ -lean_object* x_36; lean_object* x_37; -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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); -return x_37; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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) { -_start: -{ -lean_object* x_7; lean_object* x_8; 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; uint8_t x_25; -x_7 = lean_ctor_get(x_4, 5); -x_8 = lean_st_ref_get(x_5, 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 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_st_ref_get(x_5, x_10); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_st_ref_get(x_3, x_13); -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(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_4, 2); -x_19 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6; -lean_inc(x_18); -x_20 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_20, 0, x_11); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_20, 2, x_17); -lean_ctor_set(x_20, 3, x_18); -x_21 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_2); -x_22 = lean_st_ref_take(x_5, x_16); -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_is_exclusive(x_23); -if (x_25 == 0) -{ -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; uint8_t x_33; -x_26 = lean_ctor_get(x_23, 3); -x_27 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -x_28 = 0; -x_29 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_21); -lean_ctor_set(x_29, 2, x_27); -lean_ctor_set_uint8(x_29, sizeof(void*)*3, x_28); -lean_inc(x_7); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_7); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Std_PersistentArray_push___rarg(x_26, x_30); -lean_ctor_set(x_23, 3, x_31); -x_32 = lean_st_ref_set(x_5, x_23, x_24); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -x_35 = lean_box(0); -lean_ctor_set(x_32, 0, x_35); -return x_32; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_32, 1); -lean_inc(x_36); -lean_dec(x_32); -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_36); -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; 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_39 = lean_ctor_get(x_23, 0); -x_40 = lean_ctor_get(x_23, 1); -x_41 = lean_ctor_get(x_23, 2); -x_42 = lean_ctor_get(x_23, 3); -x_43 = lean_ctor_get(x_23, 4); -x_44 = lean_ctor_get(x_23, 5); -x_45 = lean_ctor_get(x_23, 6); -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_23); -x_46 = l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1; -x_47 = 0; -x_48 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_48, 0, x_1); -lean_ctor_set(x_48, 1, x_21); -lean_ctor_set(x_48, 2, x_46); -lean_ctor_set_uint8(x_48, sizeof(void*)*3, x_47); -lean_inc(x_7); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_7); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Std_PersistentArray_push___rarg(x_42, x_49); -x_51 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_51, 0, x_39); -lean_ctor_set(x_51, 1, x_40); -lean_ctor_set(x_51, 2, x_41); -lean_ctor_set(x_51, 3, x_50); -lean_ctor_set(x_51, 4, x_43); -lean_ctor_set(x_51, 5, x_44); -lean_ctor_set(x_51, 6, x_45); -x_52 = lean_st_ref_set(x_5, x_51, x_24); -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); -} -x_55 = lean_box(0); -if (lean_is_scalar(x_54)) { - x_56 = lean_alloc_ctor(0, 2, 0); -} else { - x_56 = x_54; -} -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; -} -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointChecker.containsNoJp", 54); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2() { -_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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; -x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(125u); -x_4 = lean_unsigned_to_nat(39u); -x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Join point ", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" in forbidden position", 22); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; -x_6 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_1, x_2, x_3, x_4, x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -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_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2; -x_10 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_9, x_2, x_3, x_4, x_8); -return x_10; -} -else -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_6); -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_6, 1); -x_13 = lean_ctor_get(x_6, 0); -lean_dec(x_13); -x_14 = lean_ctor_get(x_7, 0); -lean_inc(x_14); -lean_dec(x_7); -x_15 = l_Lean_LocalDecl_isJp(x_14); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_14); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_16 = lean_box(0); -lean_ctor_set(x_6, 0, x_16); -return x_6; -} -else -{ -lean_object* x_17; uint8_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; -lean_free_object(x_6); -x_17 = l_Lean_LocalDecl_userName(x_14); -lean_dec(x_14); -x_18 = 1; -x_19 = l_Lean_Name_toString(x_17, x_18); -x_20 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; -x_21 = lean_string_append(x_20, x_19); -lean_dec(x_19); -x_22 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4; -x_23 = lean_string_append(x_21, x_22); -x_24 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(x_25, x_2, x_3, x_4, x_12); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_26; -} -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_6, 1); -lean_inc(x_27); -lean_dec(x_6); -x_28 = lean_ctor_get(x_7, 0); -lean_inc(x_28); -lean_dec(x_7); -x_29 = l_Lean_LocalDecl_isJp(x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_28); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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_27); -return x_31; -} -else -{ -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; -x_32 = l_Lean_LocalDecl_userName(x_28); -lean_dec(x_28); -x_33 = 1; -x_34 = l_Lean_Name_toString(x_32, x_33); -x_35 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; -x_36 = lean_string_append(x_35, x_34); -lean_dec(x_34); -x_37 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4; -x_38 = lean_string_append(x_36, x_37); -x_39 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_40, 0, x_39); -x_41 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(x_40, x_2, x_3, x_4, x_27); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_41; -} -} -} -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1), 5, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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; lean_object* x_8; -lean_dec(x_2); -x_7 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1; -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_1, x_7, x_3, x_4, x_5, x_6); -return x_8; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Compiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("step", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2; -x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Checking whether ", 17); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" contains jp", 12); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp(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; -x_6 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_6, x_2, x_3, x_4, x_5); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1), 5, 0); -x_12 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_1, x_11, x_2, x_3, x_4, 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; 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_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_expr_dbg_to_string(x_1); -x_15 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5; -x_16 = lean_string_append(x_15, x_14); -lean_dec(x_14); -x_17 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6; -x_18 = lean_string_append(x_16, x_17); -x_19 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = l_Lean_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__7(x_6, x_20, x_2, x_3, x_4, x_13); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1), 5, 0); -x_24 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5(x_1, x_23, x_2, x_3, x_4, x_22); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__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_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(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_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___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: -{ -lean_object* x_8; -x_8 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__5___lambda__2(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_Lean_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___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: -{ -lean_object* x_7; -x_7 = l_Lean_addTrace___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__7(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; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.JoinPoints.JoinPointChecker.checkJoinPoints.goTailApp", 67); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2() { -_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_Compiler_JoinPoints_forEachFVar___rarg___closed__1; -x_2 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1; -x_3 = lean_unsigned_to_nat(146u); -x_4 = lean_unsigned_to_nat(39u); -x_5 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" : ", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" is not fully applied in ", 25); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp(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_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_1, x_3, x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2; -x_11 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_10, x_3, x_4, x_5, x_9); -return x_11; -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_7); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_ctor_get(x_7, 1); -x_14 = lean_ctor_get(x_7, 0); -lean_dec(x_14); -x_15 = lean_ctor_get(x_8, 0); -lean_inc(x_15); -lean_dec(x_8); -x_16 = l_Lean_LocalDecl_isJp(x_15); -if (x_16 == 0) -{ -lean_object* x_17; -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_17 = lean_box(0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); -x_20 = l_Lean_Compiler_jpArity(x_15); -x_21 = lean_nat_dec_eq(x_19, x_20); -lean_dec(x_20); -lean_dec(x_19); -if (x_21 == 0) -{ -lean_object* x_22; uint8_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; 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_free_object(x_7); -x_22 = l_Lean_LocalDecl_userName(x_15); -x_23 = 1; -x_24 = l_Lean_Name_toString(x_22, x_23); -x_25 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; -x_26 = lean_string_append(x_25, x_24); -lean_dec(x_24); -x_27 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3; -x_28 = lean_string_append(x_26, x_27); -x_29 = l_Lean_LocalDecl_type(x_15); -lean_dec(x_15); -x_30 = lean_expr_dbg_to_string(x_29); -lean_dec(x_29); -x_31 = lean_string_append(x_28, x_30); -lean_dec(x_30); -x_32 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4; -x_33 = lean_string_append(x_31, x_32); -x_34 = lean_expr_dbg_to_string(x_2); -x_35 = lean_string_append(x_33, x_34); -lean_dec(x_34); -x_36 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5; -x_37 = lean_string_append(x_35, x_36); -x_38 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(x_39, x_3, x_4, x_5, x_13); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_40; -} -else -{ -lean_object* x_41; -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_41 = lean_box(0); -lean_ctor_set(x_7, 0, x_41); -return x_7; -} -} -} -else -{ -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_7, 1); -lean_inc(x_42); -lean_dec(x_7); -x_43 = lean_ctor_get(x_8, 0); -lean_inc(x_43); -lean_dec(x_8); -x_44 = l_Lean_LocalDecl_isJp(x_43); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; -lean_dec(x_43); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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_42); -return x_46; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_47 = lean_unsigned_to_nat(0u); -x_48 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_47); -x_49 = l_Lean_Compiler_jpArity(x_43); -x_50 = lean_nat_dec_eq(x_48, x_49); -lean_dec(x_49); -lean_dec(x_48); -if (x_50 == 0) -{ -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; -x_51 = l_Lean_LocalDecl_userName(x_43); -x_52 = 1; -x_53 = l_Lean_Name_toString(x_51, x_52); -x_54 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; -x_55 = lean_string_append(x_54, x_53); -lean_dec(x_53); -x_56 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3; -x_57 = lean_string_append(x_55, x_56); -x_58 = l_Lean_LocalDecl_type(x_43); -lean_dec(x_43); -x_59 = lean_expr_dbg_to_string(x_58); -lean_dec(x_58); -x_60 = lean_string_append(x_57, x_59); -lean_dec(x_59); -x_61 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4; -x_62 = lean_string_append(x_60, x_61); -x_63 = lean_expr_dbg_to_string(x_2); -x_64 = lean_string_append(x_62, x_63); -lean_dec(x_63); -x_65 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5; -x_66 = lean_string_append(x_64, x_65); -x_67 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_67, 0, x_66); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_67); -x_69 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4(x_68, x_3, x_4, x_5, x_42); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_69; -} -else -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_43); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_42); -return x_71; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___boxed(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_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_2); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__8(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__24(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_eq(x_3, x_4); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_11 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -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 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_13; -x_9 = x_14; -goto _start; -} -else -{ -uint8_t x_18; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_9); -return x_22; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(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) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_eq(x_5, x_6); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_7); -x_13 = lean_array_uget(x_4, x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; -x_18 = lean_usize_add(x_5, x_17); -x_5 = x_18; -x_7 = x_15; -x_11 = x_16; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) -{ -return x_14; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -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_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -lean_inc(x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) -{ -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; -} -else -{ -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; -} -else -{ -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__3(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; -} -} -} -else -{ -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; -} -else -{ -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; -} -else -{ -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__4(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; -} -} -} -else -{ -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; -} -else -{ -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; -} -else -{ -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; -} -} -} -else -{ -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -else -{ -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) -{ -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -else -{ -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} -} -} -} -else -{ -uint8_t x_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; -} -} -} -} -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) -{ -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; -} -else -{ -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; -} -else -{ -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__3(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; -} -} -} -else -{ -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; -} -else -{ -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; -} -else -{ -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__4(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_41; +lean_dec(x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_41 = lean_box(0); +lean_ctor_set(x_7, 0, x_41); +return x_7; } } } else { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_7, 1); +lean_inc(x_42); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; -} -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; -} -else -{ -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +x_43 = lean_ctor_get(x_8, 0); +lean_inc(x_43); +lean_dec(x_8); +x_44 = l_Lean_LocalDecl_isJp(x_43); +if (x_44 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; -} -else -{ -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; -} -} +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_42); +return x_46; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} -} -} -} -} -else +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_47 = lean_unsigned_to_nat(0u); +x_48 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_47); +x_49 = l_Lean_Compiler_jpArity(x_43); +x_50 = lean_nat_dec_eq(x_48, x_49); +lean_dec(x_49); +lean_dec(x_48); +if (x_50 == 0) { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); +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; +x_51 = l_Lean_LocalDecl_userName(x_43); +x_52 = 1; +x_53 = l_Lean_Name_toString(x_51, x_52); +x_54 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3; +x_55 = lean_string_append(x_54, x_53); +lean_dec(x_53); +x_56 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__3; +x_57 = lean_string_append(x_55, x_56); +x_58 = l_Lean_LocalDecl_type(x_43); +lean_dec(x_43); +x_59 = lean_expr_dbg_to_string(x_58); +lean_dec(x_58); +x_60 = lean_string_append(x_57, x_59); +lean_dec(x_59); +x_61 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__4; +x_62 = lean_string_append(x_60, x_61); +x_63 = lean_expr_dbg_to_string(x_2); +x_64 = lean_string_append(x_62, x_63); +lean_dec(x_63); +x_65 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__5; +x_66 = lean_string_append(x_64, x_65); +x_67 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_67, 0, x_66); +x_68 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_68, 0, x_67); +x_69 = l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_68, x_3, x_4, x_5, x_42); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) -{ -return x_13; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -lean_inc(x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) -{ -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +return x_69; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_70; lean_object* x_71; +lean_dec(x_43); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_70 = lean_box(0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_42); +return x_71; } -else -{ -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; } } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___boxed(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_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) +lean_object* x_7; +x_7 = l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__8(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__9(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_dec(x_46); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(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) { +_start: { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_11; lean_object* x_12; lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -else +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -return x_58; -} -} -} +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_59; -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) { -return x_42; +return x_12; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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_19; lean_object* x_20; lean_object* x_21; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) -{ -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(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) { +_start: { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(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) { +_start: { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__8(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__9(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__8(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) { +_start: { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; -} -else +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; -} -} +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) -{ -return x_13; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_22; -lean_dec(x_20); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_28; +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; } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(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) { +_start: { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(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) { +_start: { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(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) { +_start: { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; -} -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(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) { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_77; +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; } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(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) { +_start: { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_22; -lean_dec(x_20); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_28; +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; } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(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) { +_start: { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(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) { +_start: { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_dec(x_46); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) -{ -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; -} -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(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) { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_77; +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; } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(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) { +_start: { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) -{ -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__24(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_13; +return x_14; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_22; -lean_dec(x_20); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_28; +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; } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(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) { +_start: { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_dec(x_46); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(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) { +_start: { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; -} -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29(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) { +_start: { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_11; lean_object* x_12; lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_77; +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; } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(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) { +_start: { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(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) { +_start: { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; } -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) -{ -return x_13; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); lean_inc(x_1); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_22; -lean_dec(x_20); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__23(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_28; +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; } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(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) { +_start: { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__24(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(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) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(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) { +_start: { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; -} -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(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) { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__23(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_77; +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; } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(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) { +_start: { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__24(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39(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) { +_start: { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; -} -else +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; -} -} +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) -{ -return x_13; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -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; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) -{ -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; -} -else +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; -} -} +return x_14; } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; -} -else -{ -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(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) { +_start: { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -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_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_dec(x_46); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) -{ -lean_dec(x_46); -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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); -return x_50; +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; } } } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(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) { +_start: { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -else +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} -} -} +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; } else { -uint8_t x_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) { -return x_42; +return x_14; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) -{ -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(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) { +_start: { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -else +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44(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) { +_start: { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_13 = lean_array_uget(x_4, x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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 { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(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) { +_start: { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_3, x_4); +if (x_10 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_11; lean_object* x_12; lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; -} -else +x_11 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = lean_apply_5(x_1, x_11, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +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); -return x_102; -} -} +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_13; +x_9 = x_14; +goto _start; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +return x_12; } +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) -{ -return x_13; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_9); +return x_22; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_5, x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +x_13 = lean_array_uget(x_4, x_5); lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_lt(x_18, x_17); -if (x_19 == 0) +x_14 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_13, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_20; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +lean_object* x_15; lean_object* x_16; size_t x_17; size_t 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 = 1; +x_18 = lean_usize_add(x_5, x_17); +x_5 = x_18; +x_7 = x_15; +x_11 = x_16; +goto _start; +} +else { -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +return x_14; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +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_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; -} -else -{ -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -else -{ -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) +lean_object* x_9; +x_9 = l_Lean_Expr_consumeMData(x_4); +switch (lean_obj_tag(x_9)) { +case 0: { -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_10; lean_object* x_11; +lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_11 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_10, x_5, x_6, x_7, x_8); +return x_11; } -else -{ -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +case 2: { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); +x_12 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_13 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_12, x_5, x_6, x_7, x_8); return x_13; } -else +case 5: { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; -} -} +lean_object* x_14; +x_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +switch (lean_obj_tag(x_14)) { +case 0: +{ +lean_object* x_15; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_15 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +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) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_3); +lean_dec(x_1); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_17); +return x_18; } else { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(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_45; lean_object* x_46; 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; +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_ctor_get(x_16, 0); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_st_ref_get(x_7, x_19); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_get(x_5, 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_64 = lean_st_ref_get(x_7, x_25); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = lean_st_ref_take(x_5, x_65); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_67, 2); +lean_inc(x_70); +lean_dec(x_67); +x_71 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_72 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_70); +x_73 = lean_st_ref_set(x_5, x_72, x_68); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = l_Lean_Compiler_visitMatch(x_9, x_20, x_5, x_6, x_7, x_74); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_dec(x_75); +x_79 = lean_ctor_get(x_76, 0); +lean_inc(x_79); +lean_dec(x_76); +x_80 = lean_ctor_get(x_77, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_77, 1); +lean_inc(x_81); +lean_dec(x_77); +lean_inc(x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_82 = lean_apply_5(x_2, x_79, x_5, x_6, x_7, x_78); +if (lean_obj_tag(x_82) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +x_84 = lean_array_get_size(x_80); +x_85 = lean_unsigned_to_nat(0u); +x_86 = lean_nat_dec_lt(x_85, x_84); +if (x_86 == 0) +{ +lean_object* x_87; uint8_t x_88; +lean_dec(x_84); +lean_dec(x_80); +x_87 = lean_array_get_size(x_81); +x_88 = lean_nat_dec_lt(x_85, x_87); +if (x_88 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_89; +lean_dec(x_87); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_89 = lean_box(0); +x_26 = x_89; +x_27 = x_83; +goto block_44; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_90; +x_90 = lean_nat_dec_le(x_87, x_87); +if (x_90 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_91; +lean_dec(x_87); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_91 = lean_box(0); +x_26 = x_91; +x_27 = x_83; +goto block_44; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; +x_92 = 0; +x_93 = lean_usize_of_nat(x_87); +lean_dec(x_87); +x_94 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__3(x_1, x_2, x_3, x_81, x_92, x_93, x_94, x_5, x_6, x_7, x_83); +lean_dec(x_81); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; lean_object* x_97; +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_26 = x_96; +x_27 = x_97; +goto block_44; +} +else +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_95, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_95, 1); +lean_inc(x_99); +lean_dec(x_95); +x_45 = x_98; +x_46 = x_99; +goto block_63; +} } } } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) +uint8_t x_100; +x_100 = lean_nat_dec_le(x_84, x_84); +if (x_100 == 0) { -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_101; uint8_t x_102; +lean_dec(x_84); +lean_dec(x_80); +x_101 = lean_array_get_size(x_81); +x_102 = lean_nat_dec_lt(x_85, x_101); +if (x_102 == 0) +{ +lean_object* x_103; +lean_dec(x_101); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_103 = lean_box(0); +x_26 = x_103; +x_27 = x_83; +goto block_44; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_104; +x_104 = lean_nat_dec_le(x_101, x_101); +if (x_104 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_105; +lean_dec(x_101); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -else -{ -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} -} -} +x_105 = lean_box(0); +x_26 = x_105; +x_27 = x_83; +goto block_44; } else { -uint8_t x_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) +size_t x_106; size_t x_107; lean_object* x_108; lean_object* x_109; +x_106 = 0; +x_107 = lean_usize_of_nat(x_101); +lean_dec(x_101); +x_108 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_109 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__4(x_1, x_2, x_3, x_81, x_106, x_107, x_108, x_5, x_6, x_7, x_83); +lean_dec(x_81); +if (lean_obj_tag(x_109) == 0) { -return x_42; +lean_object* x_110; lean_object* x_111; +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_26 = x_110; +x_27 = x_111; +goto block_44; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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_112; lean_object* x_113; +x_112 = lean_ctor_get(x_109, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_109, 1); +lean_inc(x_113); +lean_dec(x_109); +x_45 = x_112; +x_46 = x_113; +goto block_63; } } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +size_t x_114; size_t x_115; lean_object* x_116; lean_object* x_117; +x_114 = 0; +x_115 = lean_usize_of_nat(x_84); +lean_dec(x_84); +x_116 = lean_box(0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_117 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__5(x_2, x_80, x_114, x_115, x_116, x_5, x_6, x_7, x_83); +lean_dec(x_80); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_118; lean_object* x_119; uint8_t x_120; +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +lean_dec(x_117); +x_119 = lean_array_get_size(x_81); +x_120 = lean_nat_dec_lt(x_85, x_119); +if (x_120 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_119); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_26 = x_116; +x_27 = x_118; +goto block_44; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_121; +x_121 = lean_nat_dec_le(x_119, x_119); +if (x_121 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_119); +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_26 = x_116; +x_27 = x_118; +goto block_44; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; -} -} +size_t x_122; lean_object* x_123; +x_122 = lean_usize_of_nat(x_119); +lean_dec(x_119); +lean_inc(x_7); +lean_inc(x_5); +x_123 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__6(x_1, x_2, x_3, x_81, x_114, x_122, x_116, x_5, x_6, x_7, x_118); +lean_dec(x_81); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; +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_26 = x_124; +x_27 = x_125; +goto block_44; } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_123, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_123, 1); +lean_inc(x_127); +lean_dec(x_123); +x_45 = x_126; +x_46 = x_127; +goto block_63; +} +} +} } else { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_128; lean_object* x_129; +lean_dec(x_81); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_128 = lean_ctor_get(x_117, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_117, 1); +lean_inc(x_129); +lean_dec(x_117); +x_45 = x_128; +x_46 = x_129; +goto block_63; } -else -{ -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; } } } else { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_130; lean_object* x_131; +lean_dec(x_81); +lean_dec(x_80); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_130 = lean_ctor_get(x_82, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_82, 1); +lean_inc(x_131); +lean_dec(x_82); +x_45 = x_130; +x_46 = x_131; +goto block_63; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; +block_44: +{ +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; uint8_t x_40; +x_28 = lean_st_ref_get(x_7, x_27); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_st_ref_get(x_5, 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); +lean_dec(x_30); +x_33 = lean_ctor_get(x_24, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_24, 1); +lean_inc(x_34); +lean_dec(x_24); +x_35 = lean_ctor_get(x_31, 2); +lean_inc(x_35); +lean_dec(x_31); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_34); +lean_ctor_set(x_36, 2, x_35); +x_37 = lean_st_ref_get(x_7, x_32); +lean_dec(x_7); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_st_ref_set(x_5, x_36, x_38); +lean_dec(x_5); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 0); +lean_dec(x_41); +lean_ctor_set(x_39, 0, x_26); +return x_39; } else { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_26); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +block_63: { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +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; uint8_t x_59; +x_47 = lean_st_ref_get(x_7, x_46); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_st_ref_get(x_5, x_48); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_ctor_get(x_24, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_24, 1); +lean_inc(x_53); +lean_dec(x_24); +x_54 = lean_ctor_get(x_50, 2); +lean_inc(x_54); +lean_dec(x_50); +x_55 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_st_ref_get(x_7, x_51); lean_dec(x_7); -lean_dec(x_6); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_set(x_5, x_55, x_57); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set_tag(x_58, 1); +lean_ctor_set(x_58, 0, x_45); +return x_58; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_45); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} } } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); +uint8_t x_132; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; +x_132 = !lean_is_exclusive(x_15); +if (x_132 == 0) +{ +return x_15; } +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_15, 0); +x_134 = lean_ctor_get(x_15, 1); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_15); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } +case 1: +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_3); +x_136 = lean_ctor_get(x_9, 1); +lean_inc(x_136); +x_137 = lean_ctor_get(x_14, 0); +lean_inc(x_137); +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_138 = lean_apply_6(x_1, x_137, x_9, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +lean_dec(x_138); +x_140 = lean_apply_5(x_2, x_136, x_5, x_6, x_7, x_139); +return x_140; } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_141; +lean_dec(x_136); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_141 = !lean_is_exclusive(x_138); +if (x_141 == 0) { -return x_13; +return x_138; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_138, 0); +x_143 = lean_ctor_get(x_138, 1); +lean_inc(x_143); +lean_inc(x_142); +lean_dec(x_138); +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; } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___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) { -_start: +case 2: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); +lean_object* x_145; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -lean_inc(x_1); +x_145 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; +lean_dec(x_3); +lean_dec(x_1); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_147); +return x_148; +} +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; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_175; lean_object* x_176; 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; +x_149 = lean_ctor_get(x_145, 1); +lean_inc(x_149); +lean_dec(x_145); +x_150 = lean_ctor_get(x_146, 0); +lean_inc(x_150); +lean_dec(x_146); +x_151 = lean_st_ref_get(x_7, x_149); +x_152 = lean_ctor_get(x_151, 1); +lean_inc(x_152); +lean_dec(x_151); +x_153 = lean_st_ref_get(x_5, x_152); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +x_194 = lean_st_ref_get(x_7, x_155); +x_195 = lean_ctor_get(x_194, 1); +lean_inc(x_195); +lean_dec(x_194); +x_196 = lean_st_ref_take(x_5, x_195); +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec(x_196); +x_199 = lean_ctor_get(x_197, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_197, 2); +lean_inc(x_200); +lean_dec(x_197); +x_201 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_202 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_201); +lean_ctor_set(x_202, 2, x_200); +x_203 = lean_st_ref_set(x_5, x_202, x_198); +x_204 = lean_ctor_get(x_203, 1); +lean_inc(x_204); +lean_dec(x_203); +x_205 = l_Lean_Compiler_visitMatch(x_9, x_150, x_5, x_6, x_7, x_204); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_206, 1); +lean_inc(x_207); +x_208 = lean_ctor_get(x_205, 1); +lean_inc(x_208); +lean_dec(x_205); +x_209 = lean_ctor_get(x_206, 0); +lean_inc(x_209); +lean_dec(x_206); +x_210 = lean_ctor_get(x_207, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +lean_dec(x_207); +lean_inc(x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_212 = lean_apply_5(x_2, x_209, x_5, x_6, x_7, x_208); +if (lean_obj_tag(x_212) == 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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_lt(x_18, x_17); -if (x_19 == 0) +lean_object* x_213; lean_object* x_214; lean_object* x_215; uint8_t x_216; +x_213 = lean_ctor_get(x_212, 1); +lean_inc(x_213); +lean_dec(x_212); +x_214 = lean_array_get_size(x_210); +x_215 = lean_unsigned_to_nat(0u); +x_216 = lean_nat_dec_lt(x_215, x_214); +if (x_216 == 0) { -lean_object* x_20; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +lean_object* x_217; uint8_t x_218; +lean_dec(x_214); +lean_dec(x_210); +x_217 = lean_array_get_size(x_211); +x_218 = lean_nat_dec_lt(x_215, x_217); +if (x_218 == 0) { -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_219; +lean_dec(x_217); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_219 = lean_box(0); +x_156 = x_219; +x_157 = x_213; +goto block_174; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) +uint8_t x_220; +x_220 = lean_nat_dec_le(x_217, x_217); +if (x_220 == 0) { -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_221; +lean_dec(x_217); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_221 = lean_box(0); +x_156 = x_221; +x_157 = x_213; +goto block_174; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; +size_t x_222; size_t x_223; lean_object* x_224; lean_object* x_225; +x_222 = 0; +x_223 = lean_usize_of_nat(x_217); +lean_dec(x_217); +x_224 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_225 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__7(x_1, x_2, x_3, x_211, x_222, x_223, x_224, x_5, x_6, x_7, x_213); +lean_dec(x_211); +if (lean_obj_tag(x_225) == 0) +{ +lean_object* x_226; lean_object* x_227; +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_225, 1); +lean_inc(x_227); +lean_dec(x_225); +x_156 = x_226; +x_157 = x_227; +goto block_174; +} +else +{ +lean_object* x_228; lean_object* x_229; +x_228 = lean_ctor_get(x_225, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_225, 1); +lean_inc(x_229); +lean_dec(x_225); +x_175 = x_228; +x_176 = x_229; +goto block_193; +} } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) +uint8_t x_230; +x_230 = lean_nat_dec_le(x_214, x_214); +if (x_230 == 0) { -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) +lean_object* x_231; uint8_t x_232; +lean_dec(x_214); +lean_dec(x_210); +x_231 = lean_array_get_size(x_211); +x_232 = lean_nat_dec_lt(x_215, x_231); +if (x_232 == 0) { -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_233; +lean_dec(x_231); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_233 = lean_box(0); +x_156 = x_233; +x_157 = x_213; +goto block_174; } else { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +uint8_t x_234; +x_234 = lean_nat_dec_le(x_231, x_231); +if (x_234 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_235; +lean_dec(x_231); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_235 = lean_box(0); +x_156 = x_235; +x_157 = x_213; +goto block_174; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +size_t x_236; size_t x_237; lean_object* x_238; lean_object* x_239; +x_236 = 0; +x_237 = lean_usize_of_nat(x_231); +lean_dec(x_231); +x_238 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_239 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__8(x_1, x_2, x_3, x_211, x_236, x_237, x_238, x_5, x_6, x_7, x_213); +lean_dec(x_211); +if (lean_obj_tag(x_239) == 0) +{ +lean_object* x_240; lean_object* x_241; +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_156 = x_240; +x_157 = x_241; +goto block_174; +} +else +{ +lean_object* x_242; lean_object* x_243; +x_242 = lean_ctor_get(x_239, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_239, 1); +lean_inc(x_243); +lean_dec(x_239); +x_175 = x_242; +x_176 = x_243; +goto block_193; +} } } } else { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); +size_t x_244; size_t x_245; lean_object* x_246; lean_object* x_247; +x_244 = 0; +x_245 = lean_usize_of_nat(x_214); +lean_dec(x_214); +x_246 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_inc(x_2); +x_247 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__9(x_2, x_210, x_244, x_245, x_246, x_5, x_6, x_7, x_213); +lean_dec(x_210); +if (lean_obj_tag(x_247) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +lean_object* x_248; lean_object* x_249; uint8_t x_250; +x_248 = lean_ctor_get(x_247, 1); +lean_inc(x_248); +lean_dec(x_247); +x_249 = lean_array_get_size(x_211); +x_250 = lean_nat_dec_lt(x_215, x_249); +if (x_250 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_249); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_156 = x_246; +x_157 = x_248; +goto block_174; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_251; +x_251 = lean_nat_dec_le(x_249, x_249); +if (x_251 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_249); +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_156 = x_246; +x_157 = x_248; +goto block_174; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +size_t x_252; lean_object* x_253; +x_252 = lean_usize_of_nat(x_249); +lean_dec(x_249); +lean_inc(x_7); +lean_inc(x_5); +x_253 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__10(x_1, x_2, x_3, x_211, x_244, x_252, x_246, x_5, x_6, x_7, x_248); +lean_dec(x_211); +if (lean_obj_tag(x_253) == 0) +{ +lean_object* x_254; lean_object* x_255; +x_254 = lean_ctor_get(x_253, 0); +lean_inc(x_254); +x_255 = lean_ctor_get(x_253, 1); +lean_inc(x_255); +lean_dec(x_253); +x_156 = x_254; +x_157 = x_255; +goto block_174; +} +else +{ +lean_object* x_256; lean_object* x_257; +x_256 = lean_ctor_get(x_253, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_253, 1); +lean_inc(x_257); +lean_dec(x_253); +x_175 = x_256; +x_176 = x_257; +goto block_193; +} } } } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_258; lean_object* x_259; +lean_dec(x_211); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_258 = lean_ctor_get(x_247, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_247, 1); +lean_inc(x_259); +lean_dec(x_247); +x_175 = x_258; +x_176 = x_259; +goto block_193; +} +} +} } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) -{ -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_260; lean_object* x_261; +lean_dec(x_211); +lean_dec(x_210); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_260 = lean_ctor_get(x_212, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_212, 1); +lean_inc(x_261); +lean_dec(x_212); +x_175 = x_260; +x_176 = x_261; +goto block_193; +} +block_174: +{ +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; uint8_t x_170; +x_158 = lean_st_ref_get(x_7, x_157); +x_159 = lean_ctor_get(x_158, 1); +lean_inc(x_159); +lean_dec(x_158); +x_160 = lean_st_ref_get(x_5, x_159); +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = lean_ctor_get(x_154, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_154, 1); +lean_inc(x_164); +lean_dec(x_154); +x_165 = lean_ctor_get(x_161, 2); +lean_inc(x_165); +lean_dec(x_161); +x_166 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_166, 0, x_163); +lean_ctor_set(x_166, 1, x_164); +lean_ctor_set(x_166, 2, x_165); +x_167 = lean_st_ref_get(x_7, x_162); +lean_dec(x_7); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +lean_dec(x_167); +x_169 = lean_st_ref_set(x_5, x_166, x_168); +lean_dec(x_5); +x_170 = !lean_is_exclusive(x_169); +if (x_170 == 0) +{ +lean_object* x_171; +x_171 = lean_ctor_get(x_169, 0); +lean_dec(x_171); +lean_ctor_set(x_169, 0, x_156); +return x_169; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; +lean_object* x_172; lean_object* x_173; +x_172 = lean_ctor_get(x_169, 1); +lean_inc(x_172); +lean_dec(x_169); +x_173 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_173, 0, x_156); +lean_ctor_set(x_173, 1, x_172); +return x_173; +} +} +block_193: +{ +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; +x_177 = lean_st_ref_get(x_7, x_176); +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +lean_dec(x_177); +x_179 = lean_st_ref_get(x_5, x_178); +x_180 = lean_ctor_get(x_179, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_179, 1); +lean_inc(x_181); +lean_dec(x_179); +x_182 = lean_ctor_get(x_154, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_154, 1); +lean_inc(x_183); +lean_dec(x_154); +x_184 = lean_ctor_get(x_180, 2); +lean_inc(x_184); +lean_dec(x_180); +x_185 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +lean_ctor_set(x_185, 2, x_184); +x_186 = lean_st_ref_get(x_7, x_181); +lean_dec(x_7); +x_187 = lean_ctor_get(x_186, 1); +lean_inc(x_187); +lean_dec(x_186); +x_188 = lean_st_ref_set(x_5, x_185, x_187); +lean_dec(x_5); +x_189 = !lean_is_exclusive(x_188); +if (x_189 == 0) +{ +lean_object* x_190; +x_190 = lean_ctor_get(x_188, 0); +lean_dec(x_190); +lean_ctor_set_tag(x_188, 1); +lean_ctor_set(x_188, 0, x_175); +return x_188; +} +else +{ +lean_object* x_191; lean_object* x_192; +x_191 = lean_ctor_get(x_188, 1); +lean_inc(x_191); +lean_dec(x_188); +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_175); +lean_ctor_set(x_192, 1, x_191); +return x_192; } } } } else { -uint8_t x_59; -lean_dec(x_12); +uint8_t x_262; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) +x_262 = !lean_is_exclusive(x_145); +if (x_262 == 0) { -return x_42; +return x_145; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_145, 0); +x_264 = lean_ctor_get(x_145, 1); +lean_inc(x_264); +lean_inc(x_263); +lean_dec(x_145); +x_265 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +return x_265; } } } -else +case 3: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +lean_object* x_266; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_266 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_266) == 0) +{ +lean_object* x_267; +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +if (lean_obj_tag(x_267) == 0) +{ +lean_object* x_268; lean_object* x_269; +lean_dec(x_3); +lean_dec(x_1); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_268); +return x_269; +} +else +{ +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_296; lean_object* x_297; 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; +x_270 = lean_ctor_get(x_266, 1); +lean_inc(x_270); +lean_dec(x_266); +x_271 = lean_ctor_get(x_267, 0); +lean_inc(x_271); +lean_dec(x_267); +x_272 = lean_st_ref_get(x_7, x_270); +x_273 = lean_ctor_get(x_272, 1); +lean_inc(x_273); +lean_dec(x_272); +x_274 = lean_st_ref_get(x_5, x_273); +x_275 = lean_ctor_get(x_274, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_274, 1); +lean_inc(x_276); +lean_dec(x_274); +x_315 = lean_st_ref_get(x_7, x_276); +x_316 = lean_ctor_get(x_315, 1); +lean_inc(x_316); +lean_dec(x_315); +x_317 = lean_st_ref_take(x_5, x_316); +x_318 = lean_ctor_get(x_317, 0); +lean_inc(x_318); +x_319 = lean_ctor_get(x_317, 1); +lean_inc(x_319); +lean_dec(x_317); +x_320 = lean_ctor_get(x_318, 0); +lean_inc(x_320); +x_321 = lean_ctor_get(x_318, 2); +lean_inc(x_321); +lean_dec(x_318); +x_322 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_323 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_323, 0, x_320); +lean_ctor_set(x_323, 1, x_322); +lean_ctor_set(x_323, 2, x_321); +x_324 = lean_st_ref_set(x_5, x_323, x_319); +x_325 = lean_ctor_get(x_324, 1); +lean_inc(x_325); +lean_dec(x_324); +x_326 = l_Lean_Compiler_visitMatch(x_9, x_271, x_5, x_6, x_7, x_325); +x_327 = lean_ctor_get(x_326, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_327, 1); +lean_inc(x_328); +x_329 = lean_ctor_get(x_326, 1); +lean_inc(x_329); +lean_dec(x_326); +x_330 = lean_ctor_get(x_327, 0); +lean_inc(x_330); +lean_dec(x_327); +x_331 = lean_ctor_get(x_328, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_328, 1); +lean_inc(x_332); +lean_dec(x_328); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_333 = lean_apply_5(x_2, x_330, x_5, x_6, x_7, x_329); +if (lean_obj_tag(x_333) == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_334; lean_object* x_335; lean_object* x_336; uint8_t x_337; +x_334 = lean_ctor_get(x_333, 1); +lean_inc(x_334); +lean_dec(x_333); +x_335 = lean_array_get_size(x_331); +x_336 = lean_unsigned_to_nat(0u); +x_337 = lean_nat_dec_lt(x_336, x_335); +if (x_337 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_338; uint8_t x_339; +lean_dec(x_335); +lean_dec(x_331); +x_338 = lean_array_get_size(x_332); +x_339 = lean_nat_dec_lt(x_336, x_338); +if (x_339 == 0) +{ +lean_object* x_340; +lean_dec(x_338); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_340 = lean_box(0); +x_277 = x_340; +x_278 = x_334; +goto block_295; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_341; +x_341 = lean_nat_dec_le(x_338, x_338); +if (x_341 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_342; +lean_dec(x_338); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_342 = lean_box(0); +x_277 = x_342; +x_278 = x_334; +goto block_295; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; +size_t x_343; size_t x_344; lean_object* x_345; lean_object* x_346; +x_343 = 0; +x_344 = lean_usize_of_nat(x_338); +lean_dec(x_338); +x_345 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_346 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__11(x_1, x_2, x_3, x_332, x_343, x_344, x_345, x_5, x_6, x_7, x_334); +lean_dec(x_332); +if (lean_obj_tag(x_346) == 0) +{ +lean_object* x_347; lean_object* x_348; +x_347 = lean_ctor_get(x_346, 0); +lean_inc(x_347); +x_348 = lean_ctor_get(x_346, 1); +lean_inc(x_348); +lean_dec(x_346); +x_277 = x_347; +x_278 = x_348; +goto block_295; +} +else +{ +lean_object* x_349; lean_object* x_350; +x_349 = lean_ctor_get(x_346, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_346, 1); +lean_inc(x_350); +lean_dec(x_346); +x_296 = x_349; +x_297 = x_350; +goto block_314; +} } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) +uint8_t x_351; +x_351 = lean_nat_dec_le(x_335, x_335); +if (x_351 == 0) { -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) +lean_object* x_352; uint8_t x_353; +lean_dec(x_335); +lean_dec(x_331); +x_352 = lean_array_get_size(x_332); +x_353 = lean_nat_dec_lt(x_336, x_352); +if (x_353 == 0) { -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_354; +lean_dec(x_352); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_354 = lean_box(0); +x_277 = x_354; +x_278 = x_334; +goto block_295; } else { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_355; +x_355 = lean_nat_dec_le(x_352, x_352); +if (x_355 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_356; +lean_dec(x_352); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_356 = lean_box(0); +x_277 = x_356; +x_278 = x_334; +goto block_295; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +size_t x_357; size_t x_358; lean_object* x_359; lean_object* x_360; +x_357 = 0; +x_358 = lean_usize_of_nat(x_352); +lean_dec(x_352); +x_359 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_360 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__12(x_1, x_2, x_3, x_332, x_357, x_358, x_359, x_5, x_6, x_7, x_334); +lean_dec(x_332); +if (lean_obj_tag(x_360) == 0) +{ +lean_object* x_361; lean_object* x_362; +x_361 = lean_ctor_get(x_360, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_360, 1); +lean_inc(x_362); +lean_dec(x_360); +x_277 = x_361; +x_278 = x_362; +goto block_295; +} +else +{ +lean_object* x_363; lean_object* x_364; +x_363 = lean_ctor_get(x_360, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_360, 1); +lean_inc(x_364); +lean_dec(x_360); +x_296 = x_363; +x_297 = x_364; +goto block_314; +} } } } else { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); +size_t x_365; size_t x_366; lean_object* x_367; lean_object* x_368; +x_365 = 0; +x_366 = lean_usize_of_nat(x_335); +lean_dec(x_335); +x_367 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) +lean_inc(x_2); +x_368 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__13(x_2, x_331, x_365, x_366, x_367, x_5, x_6, x_7, x_334); +lean_dec(x_331); +if (lean_obj_tag(x_368) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) +lean_object* x_369; lean_object* x_370; uint8_t x_371; +x_369 = lean_ctor_get(x_368, 1); +lean_inc(x_369); +lean_dec(x_368); +x_370 = lean_array_get_size(x_332); +x_371 = lean_nat_dec_lt(x_336, x_370); +if (x_371 == 0) { -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_370); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; -} -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; +x_277 = x_367; +x_278 = x_369; +goto block_295; } else { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +uint8_t x_372; +x_372 = lean_nat_dec_le(x_370, x_370); +if (x_372 == 0) { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_370); +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; +x_277 = x_367; +x_278 = x_369; +goto block_295; } -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +else +{ +size_t x_373; lean_object* x_374; +x_373 = lean_usize_of_nat(x_370); +lean_dec(x_370); +lean_inc(x_7); +lean_inc(x_5); +x_374 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__14(x_1, x_2, x_3, x_332, x_365, x_373, x_367, x_5, x_6, x_7, x_369); +lean_dec(x_332); +if (lean_obj_tag(x_374) == 0) +{ +lean_object* x_375; lean_object* x_376; +x_375 = lean_ctor_get(x_374, 0); +lean_inc(x_375); +x_376 = lean_ctor_get(x_374, 1); +lean_inc(x_376); +lean_dec(x_374); +x_277 = x_375; +x_278 = x_376; +goto block_295; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_377; lean_object* x_378; +x_377 = lean_ctor_get(x_374, 0); +lean_inc(x_377); +x_378 = lean_ctor_get(x_374, 1); +lean_inc(x_378); +lean_dec(x_374); +x_296 = x_377; +x_297 = x_378; +goto block_314; +} } } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_379; lean_object* x_380; +lean_dec(x_332); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; +x_379 = lean_ctor_get(x_368, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_368, 1); +lean_inc(x_380); +lean_dec(x_368); +x_296 = x_379; +x_297 = x_380; +goto block_314; +} } } } +else +{ +lean_object* x_381; lean_object* x_382; +lean_dec(x_332); +lean_dec(x_331); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_381 = lean_ctor_get(x_333, 0); +lean_inc(x_381); +x_382 = lean_ctor_get(x_333, 1); +lean_inc(x_382); +lean_dec(x_333); +x_296 = x_381; +x_297 = x_382; +goto block_314; +} +block_295: +{ +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; uint8_t x_291; +x_279 = lean_st_ref_get(x_7, x_278); +x_280 = lean_ctor_get(x_279, 1); +lean_inc(x_280); +lean_dec(x_279); +x_281 = lean_st_ref_get(x_5, x_280); +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_ctor_get(x_275, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_275, 1); +lean_inc(x_285); +lean_dec(x_275); +x_286 = lean_ctor_get(x_282, 2); +lean_inc(x_286); +lean_dec(x_282); +x_287 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_287, 0, x_284); +lean_ctor_set(x_287, 1, x_285); +lean_ctor_set(x_287, 2, x_286); +x_288 = lean_st_ref_get(x_7, x_283); +lean_dec(x_7); +x_289 = lean_ctor_get(x_288, 1); +lean_inc(x_289); +lean_dec(x_288); +x_290 = lean_st_ref_set(x_5, x_287, x_289); +lean_dec(x_5); +x_291 = !lean_is_exclusive(x_290); +if (x_291 == 0) +{ +lean_object* x_292; +x_292 = lean_ctor_get(x_290, 0); +lean_dec(x_292); +lean_ctor_set(x_290, 0, x_277); +return x_290; +} +else +{ +lean_object* x_293; lean_object* x_294; +x_293 = lean_ctor_get(x_290, 1); +lean_inc(x_293); +lean_dec(x_290); +x_294 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_294, 0, x_277); +lean_ctor_set(x_294, 1, x_293); +return x_294; } } -else +block_314: { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); +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; uint8_t x_310; +x_298 = lean_st_ref_get(x_7, x_297); +x_299 = lean_ctor_get(x_298, 1); +lean_inc(x_299); +lean_dec(x_298); +x_300 = lean_st_ref_get(x_5, x_299); +x_301 = lean_ctor_get(x_300, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_300, 1); +lean_inc(x_302); +lean_dec(x_300); +x_303 = lean_ctor_get(x_275, 0); +lean_inc(x_303); +x_304 = lean_ctor_get(x_275, 1); +lean_inc(x_304); +lean_dec(x_275); +x_305 = lean_ctor_get(x_301, 2); +lean_inc(x_305); +lean_dec(x_301); +x_306 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_306, 0, x_303); +lean_ctor_set(x_306, 1, x_304); +lean_ctor_set(x_306, 2, x_305); +x_307 = lean_st_ref_get(x_7, x_302); +lean_dec(x_7); +x_308 = lean_ctor_get(x_307, 1); +lean_inc(x_308); +lean_dec(x_307); +x_309 = lean_st_ref_set(x_5, x_306, x_308); +lean_dec(x_5); +x_310 = !lean_is_exclusive(x_309); +if (x_310 == 0) +{ +lean_object* x_311; +x_311 = lean_ctor_get(x_309, 0); +lean_dec(x_311); +lean_ctor_set_tag(x_309, 1); +lean_ctor_set(x_309, 0, x_296); +return x_309; +} +else +{ +lean_object* x_312; lean_object* x_313; +x_312 = lean_ctor_get(x_309, 1); +lean_inc(x_312); +lean_dec(x_309); +x_313 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_313, 0, x_296); +lean_ctor_set(x_313, 1, x_312); +return x_313; +} +} +} +} +else +{ +uint8_t x_383; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_383 = !lean_is_exclusive(x_266); +if (x_383 == 0) { -return x_13; +return x_266; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_384; lean_object* x_385; lean_object* x_386; +x_384 = lean_ctor_get(x_266, 0); +x_385 = lean_ctor_get(x_266, 1); +lean_inc(x_385); +lean_inc(x_384); +lean_dec(x_266); +x_386 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_386, 0, x_384); +lean_ctor_set(x_386, 1, x_385); +return x_386; } } } -} -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__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: +case 4: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); +lean_object* x_387; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -lean_inc(x_1); +x_387 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_387) == 0) +{ +lean_object* x_388; +x_388 = lean_ctor_get(x_387, 0); +lean_inc(x_388); +if (lean_obj_tag(x_388) == 0) +{ +lean_object* x_389; lean_object* x_390; +lean_dec(x_3); +lean_dec(x_1); +x_389 = lean_ctor_get(x_387, 1); +lean_inc(x_389); +lean_dec(x_387); +x_390 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_389); +return x_390; +} +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; lean_object* x_417; lean_object* x_418; 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; +x_391 = lean_ctor_get(x_387, 1); +lean_inc(x_391); +lean_dec(x_387); +x_392 = lean_ctor_get(x_388, 0); +lean_inc(x_392); +lean_dec(x_388); +x_393 = lean_st_ref_get(x_7, x_391); +x_394 = lean_ctor_get(x_393, 1); +lean_inc(x_394); +lean_dec(x_393); +x_395 = lean_st_ref_get(x_5, x_394); +x_396 = lean_ctor_get(x_395, 0); +lean_inc(x_396); +x_397 = lean_ctor_get(x_395, 1); +lean_inc(x_397); +lean_dec(x_395); +x_436 = lean_st_ref_get(x_7, x_397); +x_437 = lean_ctor_get(x_436, 1); +lean_inc(x_437); +lean_dec(x_436); +x_438 = lean_st_ref_take(x_5, x_437); +x_439 = lean_ctor_get(x_438, 0); +lean_inc(x_439); +x_440 = lean_ctor_get(x_438, 1); +lean_inc(x_440); +lean_dec(x_438); +x_441 = lean_ctor_get(x_439, 0); +lean_inc(x_441); +x_442 = lean_ctor_get(x_439, 2); +lean_inc(x_442); +lean_dec(x_439); +x_443 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_444 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_444, 0, x_441); +lean_ctor_set(x_444, 1, x_443); +lean_ctor_set(x_444, 2, x_442); +x_445 = lean_st_ref_set(x_5, x_444, x_440); +x_446 = lean_ctor_get(x_445, 1); +lean_inc(x_446); +lean_dec(x_445); +x_447 = l_Lean_Compiler_visitMatch(x_9, x_392, x_5, x_6, x_7, x_446); +x_448 = lean_ctor_get(x_447, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_448, 1); +lean_inc(x_449); +x_450 = lean_ctor_get(x_447, 1); +lean_inc(x_450); +lean_dec(x_447); +x_451 = lean_ctor_get(x_448, 0); +lean_inc(x_451); +lean_dec(x_448); +x_452 = lean_ctor_get(x_449, 0); +lean_inc(x_452); +x_453 = lean_ctor_get(x_449, 1); +lean_inc(x_453); +lean_dec(x_449); +lean_inc(x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_454 = lean_apply_5(x_2, x_451, x_5, x_6, x_7, x_450); +if (lean_obj_tag(x_454) == 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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_lt(x_18, x_17); -if (x_19 == 0) +lean_object* x_455; lean_object* x_456; lean_object* x_457; uint8_t x_458; +x_455 = lean_ctor_get(x_454, 1); +lean_inc(x_455); +lean_dec(x_454); +x_456 = lean_array_get_size(x_452); +x_457 = lean_unsigned_to_nat(0u); +x_458 = lean_nat_dec_lt(x_457, x_456); +if (x_458 == 0) { -lean_object* x_20; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +lean_object* x_459; uint8_t x_460; +lean_dec(x_456); +lean_dec(x_452); +x_459 = lean_array_get_size(x_453); +x_460 = lean_nat_dec_lt(x_457, x_459); +if (x_460 == 0) { -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_461; +lean_dec(x_459); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_461 = lean_box(0); +x_398 = x_461; +x_399 = x_455; +goto block_416; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) +uint8_t x_462; +x_462 = lean_nat_dec_le(x_459, x_459); +if (x_462 == 0) { -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_463; +lean_dec(x_459); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_463 = lean_box(0); +x_398 = x_463; +x_399 = x_455; +goto block_416; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; +size_t x_464; size_t x_465; lean_object* x_466; lean_object* x_467; +x_464 = 0; +x_465 = lean_usize_of_nat(x_459); +lean_dec(x_459); +x_466 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_467 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__15(x_1, x_2, x_3, x_453, x_464, x_465, x_466, x_5, x_6, x_7, x_455); +lean_dec(x_453); +if (lean_obj_tag(x_467) == 0) +{ +lean_object* x_468; lean_object* x_469; +x_468 = lean_ctor_get(x_467, 0); +lean_inc(x_468); +x_469 = lean_ctor_get(x_467, 1); +lean_inc(x_469); +lean_dec(x_467); +x_398 = x_468; +x_399 = x_469; +goto block_416; +} +else +{ +lean_object* x_470; lean_object* x_471; +x_470 = lean_ctor_get(x_467, 0); +lean_inc(x_470); +x_471 = lean_ctor_get(x_467, 1); +lean_inc(x_471); +lean_dec(x_467); +x_417 = x_470; +x_418 = x_471; +goto block_435; +} } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) +uint8_t x_472; +x_472 = lean_nat_dec_le(x_456, x_456); +if (x_472 == 0) { -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) +lean_object* x_473; uint8_t x_474; +lean_dec(x_456); +lean_dec(x_452); +x_473 = lean_array_get_size(x_453); +x_474 = lean_nat_dec_lt(x_457, x_473); +if (x_474 == 0) { -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_475; +lean_dec(x_473); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_475 = lean_box(0); +x_398 = x_475; +x_399 = x_455; +goto block_416; } else { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +uint8_t x_476; +x_476 = lean_nat_dec_le(x_473, x_473); +if (x_476 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_477; +lean_dec(x_473); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_477 = lean_box(0); +x_398 = x_477; +x_399 = x_455; +goto block_416; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +size_t x_478; size_t x_479; lean_object* x_480; lean_object* x_481; +x_478 = 0; +x_479 = lean_usize_of_nat(x_473); +lean_dec(x_473); +x_480 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_481 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__16(x_1, x_2, x_3, x_453, x_478, x_479, x_480, x_5, x_6, x_7, x_455); +lean_dec(x_453); +if (lean_obj_tag(x_481) == 0) +{ +lean_object* x_482; lean_object* x_483; +x_482 = lean_ctor_get(x_481, 0); +lean_inc(x_482); +x_483 = lean_ctor_get(x_481, 1); +lean_inc(x_483); +lean_dec(x_481); +x_398 = x_482; +x_399 = x_483; +goto block_416; +} +else +{ +lean_object* x_484; lean_object* x_485; +x_484 = lean_ctor_get(x_481, 0); +lean_inc(x_484); +x_485 = lean_ctor_get(x_481, 1); +lean_inc(x_485); +lean_dec(x_481); +x_417 = x_484; +x_418 = x_485; +goto block_435; +} } } } else { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); +size_t x_486; size_t x_487; lean_object* x_488; lean_object* x_489; +x_486 = 0; +x_487 = lean_usize_of_nat(x_456); +lean_dec(x_456); +x_488 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_inc(x_2); +x_489 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__17(x_2, x_452, x_486, x_487, x_488, x_5, x_6, x_7, x_455); +lean_dec(x_452); +if (lean_obj_tag(x_489) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +lean_object* x_490; lean_object* x_491; uint8_t x_492; +x_490 = lean_ctor_get(x_489, 1); +lean_inc(x_490); +lean_dec(x_489); +x_491 = lean_array_get_size(x_453); +x_492 = lean_nat_dec_lt(x_457, x_491); +if (x_492 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_491); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_398 = x_488; +x_399 = x_490; +goto block_416; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_493; +x_493 = lean_nat_dec_le(x_491, x_491); +if (x_493 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_491); +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_398 = x_488; +x_399 = x_490; +goto block_416; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +size_t x_494; lean_object* x_495; +x_494 = lean_usize_of_nat(x_491); +lean_dec(x_491); +lean_inc(x_7); +lean_inc(x_5); +x_495 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__18(x_1, x_2, x_3, x_453, x_486, x_494, x_488, x_5, x_6, x_7, x_490); +lean_dec(x_453); +if (lean_obj_tag(x_495) == 0) +{ +lean_object* x_496; lean_object* x_497; +x_496 = lean_ctor_get(x_495, 0); +lean_inc(x_496); +x_497 = lean_ctor_get(x_495, 1); +lean_inc(x_497); +lean_dec(x_495); +x_398 = x_496; +x_399 = x_497; +goto block_416; +} +else +{ +lean_object* x_498; lean_object* x_499; +x_498 = lean_ctor_get(x_495, 0); +lean_inc(x_498); +x_499 = lean_ctor_get(x_495, 1); +lean_inc(x_499); +lean_dec(x_495); +x_417 = x_498; +x_418 = x_499; +goto block_435; +} } } } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) -{ -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_500; lean_object* x_501; +lean_dec(x_453); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_500 = lean_ctor_get(x_489, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_489, 1); +lean_inc(x_501); +lean_dec(x_489); +x_417 = x_500; +x_418 = x_501; +goto block_435; +} +} +} +} +else +{ +lean_object* x_502; lean_object* x_503; +lean_dec(x_453); +lean_dec(x_452); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_502 = lean_ctor_get(x_454, 0); +lean_inc(x_502); +x_503 = lean_ctor_get(x_454, 1); +lean_inc(x_503); +lean_dec(x_454); +x_417 = x_502; +x_418 = x_503; +goto block_435; +} +block_416: +{ +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; uint8_t x_412; +x_400 = lean_st_ref_get(x_7, x_399); +x_401 = lean_ctor_get(x_400, 1); +lean_inc(x_401); +lean_dec(x_400); +x_402 = lean_st_ref_get(x_5, x_401); +x_403 = lean_ctor_get(x_402, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_402, 1); +lean_inc(x_404); +lean_dec(x_402); +x_405 = lean_ctor_get(x_396, 0); +lean_inc(x_405); +x_406 = lean_ctor_get(x_396, 1); +lean_inc(x_406); +lean_dec(x_396); +x_407 = lean_ctor_get(x_403, 2); +lean_inc(x_407); +lean_dec(x_403); +x_408 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_408, 0, x_405); +lean_ctor_set(x_408, 1, x_406); +lean_ctor_set(x_408, 2, x_407); +x_409 = lean_st_ref_get(x_7, x_404); +lean_dec(x_7); +x_410 = lean_ctor_get(x_409, 1); +lean_inc(x_410); +lean_dec(x_409); +x_411 = lean_st_ref_set(x_5, x_408, x_410); +lean_dec(x_5); +x_412 = !lean_is_exclusive(x_411); +if (x_412 == 0) +{ +lean_object* x_413; +x_413 = lean_ctor_get(x_411, 0); +lean_dec(x_413); +lean_ctor_set(x_411, 0, x_398); +return x_411; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +lean_object* x_414; lean_object* x_415; +x_414 = lean_ctor_get(x_411, 1); +lean_inc(x_414); +lean_dec(x_411); +x_415 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_415, 0, x_398); +lean_ctor_set(x_415, 1, x_414); +return x_415; +} +} +block_435: { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); +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; uint8_t x_431; +x_419 = lean_st_ref_get(x_7, x_418); +x_420 = lean_ctor_get(x_419, 1); +lean_inc(x_420); +lean_dec(x_419); +x_421 = lean_st_ref_get(x_5, x_420); +x_422 = lean_ctor_get(x_421, 0); +lean_inc(x_422); +x_423 = lean_ctor_get(x_421, 1); +lean_inc(x_423); +lean_dec(x_421); +x_424 = lean_ctor_get(x_396, 0); +lean_inc(x_424); +x_425 = lean_ctor_get(x_396, 1); +lean_inc(x_425); +lean_dec(x_396); +x_426 = lean_ctor_get(x_422, 2); +lean_inc(x_426); +lean_dec(x_422); +x_427 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_427, 0, x_424); +lean_ctor_set(x_427, 1, x_425); +lean_ctor_set(x_427, 2, x_426); +x_428 = lean_st_ref_get(x_7, x_423); +lean_dec(x_7); +x_429 = lean_ctor_get(x_428, 1); +lean_inc(x_429); +lean_dec(x_428); +x_430 = lean_st_ref_set(x_5, x_427, x_429); +lean_dec(x_5); +x_431 = !lean_is_exclusive(x_430); +if (x_431 == 0) +{ +lean_object* x_432; +x_432 = lean_ctor_get(x_430, 0); +lean_dec(x_432); +lean_ctor_set_tag(x_430, 1); +lean_ctor_set(x_430, 0, x_417); +return x_430; +} +else +{ +lean_object* x_433; lean_object* x_434; +x_433 = lean_ctor_get(x_430, 1); +lean_inc(x_433); +lean_dec(x_430); +x_434 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_434, 0, x_417); +lean_ctor_set(x_434, 1, x_433); +return x_434; +} +} +} +} +else +{ +uint8_t x_504; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; +x_504 = !lean_is_exclusive(x_387); +if (x_504 == 0) +{ +return x_387; } else { -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; +lean_object* x_505; lean_object* x_506; lean_object* x_507; +x_505 = lean_ctor_get(x_387, 0); +x_506 = lean_ctor_get(x_387, 1); +lean_inc(x_506); +lean_inc(x_505); +lean_dec(x_387); +x_507 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_507, 0, x_505); +lean_ctor_set(x_507, 1, x_506); +return x_507; } } } +case 5: +{ +lean_object* x_508; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_508 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_508) == 0) +{ +lean_object* x_509; +x_509 = lean_ctor_get(x_508, 0); +lean_inc(x_509); +if (lean_obj_tag(x_509) == 0) +{ +lean_object* x_510; lean_object* x_511; +lean_dec(x_3); +lean_dec(x_1); +x_510 = lean_ctor_get(x_508, 1); +lean_inc(x_510); +lean_dec(x_508); +x_511 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_510); +return x_511; +} +else +{ +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_538; lean_object* x_539; 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; +x_512 = lean_ctor_get(x_508, 1); +lean_inc(x_512); +lean_dec(x_508); +x_513 = lean_ctor_get(x_509, 0); +lean_inc(x_513); +lean_dec(x_509); +x_514 = lean_st_ref_get(x_7, x_512); +x_515 = lean_ctor_get(x_514, 1); +lean_inc(x_515); +lean_dec(x_514); +x_516 = lean_st_ref_get(x_5, x_515); +x_517 = lean_ctor_get(x_516, 0); +lean_inc(x_517); +x_518 = lean_ctor_get(x_516, 1); +lean_inc(x_518); +lean_dec(x_516); +x_557 = lean_st_ref_get(x_7, x_518); +x_558 = lean_ctor_get(x_557, 1); +lean_inc(x_558); +lean_dec(x_557); +x_559 = lean_st_ref_take(x_5, x_558); +x_560 = lean_ctor_get(x_559, 0); +lean_inc(x_560); +x_561 = lean_ctor_get(x_559, 1); +lean_inc(x_561); +lean_dec(x_559); +x_562 = lean_ctor_get(x_560, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_560, 2); +lean_inc(x_563); +lean_dec(x_560); +x_564 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_565 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_565, 0, x_562); +lean_ctor_set(x_565, 1, x_564); +lean_ctor_set(x_565, 2, x_563); +x_566 = lean_st_ref_set(x_5, x_565, x_561); +x_567 = lean_ctor_get(x_566, 1); +lean_inc(x_567); +lean_dec(x_566); +x_568 = l_Lean_Compiler_visitMatch(x_9, x_513, x_5, x_6, x_7, x_567); +x_569 = lean_ctor_get(x_568, 0); +lean_inc(x_569); +x_570 = lean_ctor_get(x_569, 1); +lean_inc(x_570); +x_571 = lean_ctor_get(x_568, 1); +lean_inc(x_571); +lean_dec(x_568); +x_572 = lean_ctor_get(x_569, 0); +lean_inc(x_572); +lean_dec(x_569); +x_573 = lean_ctor_get(x_570, 0); +lean_inc(x_573); +x_574 = lean_ctor_get(x_570, 1); +lean_inc(x_574); +lean_dec(x_570); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_575 = lean_apply_5(x_2, x_572, x_5, x_6, x_7, x_571); +if (lean_obj_tag(x_575) == 0) +{ +lean_object* x_576; lean_object* x_577; lean_object* x_578; uint8_t x_579; +x_576 = lean_ctor_get(x_575, 1); +lean_inc(x_576); +lean_dec(x_575); +x_577 = lean_array_get_size(x_573); +x_578 = lean_unsigned_to_nat(0u); +x_579 = lean_nat_dec_lt(x_578, x_577); +if (x_579 == 0) +{ +lean_object* x_580; uint8_t x_581; +lean_dec(x_577); +lean_dec(x_573); +x_580 = lean_array_get_size(x_574); +x_581 = lean_nat_dec_lt(x_578, x_580); +if (x_581 == 0) +{ +lean_object* x_582; +lean_dec(x_580); +lean_dec(x_574); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_582 = lean_box(0); +x_519 = x_582; +x_520 = x_576; +goto block_537; } else { -uint8_t x_59; -lean_dec(x_12); -lean_dec(x_7); +uint8_t x_583; +x_583 = lean_nat_dec_le(x_580, x_580); +if (x_583 == 0) +{ +lean_object* x_584; +lean_dec(x_580); +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) -{ -return x_42; +x_584 = lean_box(0); +x_519 = x_584; +x_520 = x_576; +goto block_537; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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; +size_t x_585; size_t x_586; lean_object* x_587; lean_object* x_588; +x_585 = 0; +x_586 = lean_usize_of_nat(x_580); +lean_dec(x_580); +x_587 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_588 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__19(x_1, x_2, x_3, x_574, x_585, x_586, x_587, x_5, x_6, x_7, x_576); +lean_dec(x_574); +if (lean_obj_tag(x_588) == 0) +{ +lean_object* x_589; lean_object* x_590; +x_589 = lean_ctor_get(x_588, 0); +lean_inc(x_589); +x_590 = lean_ctor_get(x_588, 1); +lean_inc(x_590); +lean_dec(x_588); +x_519 = x_589; +x_520 = x_590; +goto block_537; } +else +{ +lean_object* x_591; lean_object* x_592; +x_591 = lean_ctor_get(x_588, 0); +lean_inc(x_591); +x_592 = lean_ctor_get(x_588, 1); +lean_inc(x_592); +lean_dec(x_588); +x_538 = x_591; +x_539 = x_592; +goto block_556; } } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +uint8_t x_593; +x_593 = lean_nat_dec_le(x_577, x_577); +if (x_593 == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_594; uint8_t x_595; +lean_dec(x_577); +lean_dec(x_573); +x_594 = lean_array_get_size(x_574); +x_595 = lean_nat_dec_lt(x_578, x_594); +if (x_595 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_596; +lean_dec(x_594); +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_596 = lean_box(0); +x_519 = x_596; +x_520 = x_576; +goto block_537; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_597; +x_597 = lean_nat_dec_le(x_594, x_594); +if (x_597 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_598; +lean_dec(x_594); +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_598 = lean_box(0); +x_519 = x_598; +x_520 = x_576; +goto block_537; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; +size_t x_599; size_t x_600; lean_object* x_601; lean_object* x_602; +x_599 = 0; +x_600 = lean_usize_of_nat(x_594); +lean_dec(x_594); +x_601 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_602 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__20(x_1, x_2, x_3, x_574, x_599, x_600, x_601, x_5, x_6, x_7, x_576); +lean_dec(x_574); +if (lean_obj_tag(x_602) == 0) +{ +lean_object* x_603; lean_object* x_604; +x_603 = lean_ctor_get(x_602, 0); +lean_inc(x_603); +x_604 = lean_ctor_get(x_602, 1); +lean_inc(x_604); +lean_dec(x_602); +x_519 = x_603; +x_520 = x_604; +goto block_537; +} +else +{ +lean_object* x_605; lean_object* x_606; +x_605 = lean_ctor_get(x_602, 0); +lean_inc(x_605); +x_606 = lean_ctor_get(x_602, 1); +lean_inc(x_606); +lean_dec(x_602); +x_538 = x_605; +x_539 = x_606; +goto block_556; +} } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) +size_t x_607; size_t x_608; lean_object* x_609; lean_object* x_610; +x_607 = 0; +x_608 = lean_usize_of_nat(x_577); +lean_dec(x_577); +x_609 = lean_box(0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_610 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__21(x_2, x_573, x_607, x_608, x_609, x_5, x_6, x_7, x_576); +lean_dec(x_573); +if (lean_obj_tag(x_610) == 0) { -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) +lean_object* x_611; lean_object* x_612; uint8_t x_613; +x_611 = lean_ctor_get(x_610, 1); +lean_inc(x_611); +lean_dec(x_610); +x_612 = lean_array_get_size(x_574); +x_613 = lean_nat_dec_lt(x_578, x_612); +if (x_613 == 0) { -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_612); +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_519 = x_609; +x_520 = x_611; +goto block_537; } else { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +uint8_t x_614; +x_614 = lean_nat_dec_le(x_612, x_612); +if (x_614 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_612); +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; -} -else -{ -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; -} -} +x_519 = x_609; +x_520 = x_611; +goto block_537; } else { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); +size_t x_615; lean_object* x_616; +x_615 = lean_usize_of_nat(x_612); +lean_dec(x_612); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) +x_616 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__22(x_1, x_2, x_3, x_574, x_607, x_615, x_609, x_5, x_6, x_7, x_611); +lean_dec(x_574); +if (lean_obj_tag(x_616) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); +lean_object* x_617; lean_object* x_618; +x_617 = lean_ctor_get(x_616, 0); +lean_inc(x_617); +x_618 = lean_ctor_get(x_616, 1); +lean_inc(x_618); +lean_dec(x_616); +x_519 = x_617; +x_520 = x_618; +goto block_537; } -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) +else { -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_619; lean_object* x_620; +x_619 = lean_ctor_get(x_616, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_616, 1); +lean_inc(x_620); +lean_dec(x_616); +x_538 = x_619; +x_539 = x_620; +goto block_556; +} +} +} +} +else +{ +lean_object* x_621; lean_object* x_622; +lean_dec(x_574); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; +x_621 = lean_ctor_get(x_610, 0); +lean_inc(x_621); +x_622 = lean_ctor_get(x_610, 1); +lean_inc(x_622); +lean_dec(x_610); +x_538 = x_621; +x_539 = x_622; +goto block_556; +} +} +} +} +else +{ +lean_object* x_623; lean_object* x_624; +lean_dec(x_574); +lean_dec(x_573); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_623 = lean_ctor_get(x_575, 0); +lean_inc(x_623); +x_624 = lean_ctor_get(x_575, 1); +lean_inc(x_624); +lean_dec(x_575); +x_538 = x_623; +x_539 = x_624; +goto block_556; } -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; +block_537: +{ +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; uint8_t x_533; +x_521 = lean_st_ref_get(x_7, x_520); +x_522 = lean_ctor_get(x_521, 1); +lean_inc(x_522); +lean_dec(x_521); +x_523 = lean_st_ref_get(x_5, x_522); +x_524 = lean_ctor_get(x_523, 0); +lean_inc(x_524); +x_525 = lean_ctor_get(x_523, 1); +lean_inc(x_525); +lean_dec(x_523); +x_526 = lean_ctor_get(x_517, 0); +lean_inc(x_526); +x_527 = lean_ctor_get(x_517, 1); +lean_inc(x_527); +lean_dec(x_517); +x_528 = lean_ctor_get(x_524, 2); +lean_inc(x_528); +lean_dec(x_524); +x_529 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_529, 0, x_526); +lean_ctor_set(x_529, 1, x_527); +lean_ctor_set(x_529, 2, x_528); +x_530 = lean_st_ref_get(x_7, x_525); +lean_dec(x_7); +x_531 = lean_ctor_get(x_530, 1); +lean_inc(x_531); +lean_dec(x_530); +x_532 = lean_st_ref_set(x_5, x_529, x_531); +lean_dec(x_5); +x_533 = !lean_is_exclusive(x_532); +if (x_533 == 0) +{ +lean_object* x_534; +x_534 = lean_ctor_get(x_532, 0); +lean_dec(x_534); +lean_ctor_set(x_532, 0, x_519); +return x_532; } else { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) +lean_object* x_535; lean_object* x_536; +x_535 = lean_ctor_get(x_532, 1); +lean_inc(x_535); +lean_dec(x_532); +x_536 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_536, 0, x_519); +lean_ctor_set(x_536, 1, x_535); +return x_536; +} +} +block_556: { -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); +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; uint8_t x_552; +x_540 = lean_st_ref_get(x_7, x_539); +x_541 = lean_ctor_get(x_540, 1); +lean_inc(x_541); +lean_dec(x_540); +x_542 = lean_st_ref_get(x_5, x_541); +x_543 = lean_ctor_get(x_542, 0); +lean_inc(x_543); +x_544 = lean_ctor_get(x_542, 1); +lean_inc(x_544); +lean_dec(x_542); +x_545 = lean_ctor_get(x_517, 0); +lean_inc(x_545); +x_546 = lean_ctor_get(x_517, 1); +lean_inc(x_546); +lean_dec(x_517); +x_547 = lean_ctor_get(x_543, 2); +lean_inc(x_547); +lean_dec(x_543); +x_548 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_548, 0, x_545); +lean_ctor_set(x_548, 1, x_546); +lean_ctor_set(x_548, 2, x_547); +x_549 = lean_st_ref_get(x_7, x_544); +lean_dec(x_7); +x_550 = lean_ctor_get(x_549, 1); +lean_inc(x_550); +lean_dec(x_549); +x_551 = lean_st_ref_set(x_5, x_548, x_550); +lean_dec(x_5); +x_552 = !lean_is_exclusive(x_551); +if (x_552 == 0) +{ +lean_object* x_553; +x_553 = lean_ctor_get(x_551, 0); +lean_dec(x_553); +lean_ctor_set_tag(x_551, 1); +lean_ctor_set(x_551, 0, x_538); +return x_551; +} +else +{ +lean_object* x_554; lean_object* x_555; +x_554 = lean_ctor_get(x_551, 1); +lean_inc(x_554); +lean_dec(x_551); +x_555 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_555, 0, x_538); +lean_ctor_set(x_555, 1, x_554); +return x_555; +} +} +} +} +else +{ +uint8_t x_625; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; -} -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +x_625 = !lean_is_exclusive(x_508); +if (x_625 == 0) +{ +return x_508; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_626; lean_object* x_627; lean_object* x_628; +x_626 = lean_ctor_get(x_508, 0); +x_627 = lean_ctor_get(x_508, 1); +lean_inc(x_627); +lean_inc(x_626); +lean_dec(x_508); +x_628 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_628, 0, x_626); +lean_ctor_set(x_628, 1, x_627); +return x_628; } } } +case 6: +{ +lean_object* x_629; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_629 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_629) == 0) +{ +lean_object* x_630; +x_630 = lean_ctor_get(x_629, 0); +lean_inc(x_630); +if (lean_obj_tag(x_630) == 0) +{ +lean_object* x_631; lean_object* x_632; +lean_dec(x_3); +lean_dec(x_1); +x_631 = lean_ctor_get(x_629, 1); +lean_inc(x_631); +lean_dec(x_629); +x_632 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, 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_659; lean_object* x_660; 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; +x_633 = lean_ctor_get(x_629, 1); +lean_inc(x_633); +lean_dec(x_629); +x_634 = lean_ctor_get(x_630, 0); +lean_inc(x_634); +lean_dec(x_630); +x_635 = lean_st_ref_get(x_7, x_633); +x_636 = lean_ctor_get(x_635, 1); +lean_inc(x_636); +lean_dec(x_635); +x_637 = lean_st_ref_get(x_5, x_636); +x_638 = lean_ctor_get(x_637, 0); +lean_inc(x_638); +x_639 = lean_ctor_get(x_637, 1); +lean_inc(x_639); +lean_dec(x_637); +x_678 = lean_st_ref_get(x_7, x_639); +x_679 = lean_ctor_get(x_678, 1); +lean_inc(x_679); +lean_dec(x_678); +x_680 = lean_st_ref_take(x_5, x_679); +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 = lean_ctor_get(x_681, 0); +lean_inc(x_683); +x_684 = lean_ctor_get(x_681, 2); +lean_inc(x_684); +lean_dec(x_681); +x_685 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_686 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_686, 0, x_683); +lean_ctor_set(x_686, 1, x_685); +lean_ctor_set(x_686, 2, x_684); +x_687 = lean_st_ref_set(x_5, x_686, x_682); +x_688 = lean_ctor_get(x_687, 1); +lean_inc(x_688); +lean_dec(x_687); +x_689 = l_Lean_Compiler_visitMatch(x_9, x_634, x_5, x_6, x_7, x_688); +x_690 = lean_ctor_get(x_689, 0); +lean_inc(x_690); +x_691 = lean_ctor_get(x_690, 1); +lean_inc(x_691); +x_692 = lean_ctor_get(x_689, 1); +lean_inc(x_692); +lean_dec(x_689); +x_693 = lean_ctor_get(x_690, 0); +lean_inc(x_693); +lean_dec(x_690); +x_694 = lean_ctor_get(x_691, 0); +lean_inc(x_694); +x_695 = lean_ctor_get(x_691, 1); +lean_inc(x_695); +lean_dec(x_691); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_696 = lean_apply_5(x_2, x_693, x_5, x_6, x_7, x_692); +if (lean_obj_tag(x_696) == 0) +{ +lean_object* x_697; lean_object* x_698; lean_object* x_699; uint8_t x_700; +x_697 = lean_ctor_get(x_696, 1); +lean_inc(x_697); +lean_dec(x_696); +x_698 = lean_array_get_size(x_694); +x_699 = lean_unsigned_to_nat(0u); +x_700 = lean_nat_dec_lt(x_699, x_698); +if (x_700 == 0) +{ +lean_object* x_701; uint8_t x_702; +lean_dec(x_698); +lean_dec(x_694); +x_701 = lean_array_get_size(x_695); +x_702 = lean_nat_dec_lt(x_699, x_701); +if (x_702 == 0) +{ +lean_object* x_703; +lean_dec(x_701); +lean_dec(x_695); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_703 = lean_box(0); +x_640 = x_703; +x_641 = x_697; +goto block_658; +} else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); +uint8_t x_704; +x_704 = lean_nat_dec_le(x_701, x_701); +if (x_704 == 0) +{ +lean_object* x_705; +lean_dec(x_701); +lean_dec(x_695); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); +x_705 = lean_box(0); +x_640 = x_705; +x_641 = x_697; +goto block_658; } -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; +else +{ +size_t x_706; size_t x_707; lean_object* x_708; lean_object* x_709; +x_706 = 0; +x_707 = lean_usize_of_nat(x_701); +lean_dec(x_701); +x_708 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_709 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__23(x_1, x_2, x_3, x_695, x_706, x_707, x_708, x_5, x_6, x_7, x_697); +lean_dec(x_695); +if (lean_obj_tag(x_709) == 0) +{ +lean_object* x_710; lean_object* x_711; +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_640 = x_710; +x_641 = x_711; +goto block_658; } -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; +else +{ +lean_object* x_712; lean_object* x_713; +x_712 = lean_ctor_get(x_709, 0); +lean_inc(x_712); +x_713 = lean_ctor_get(x_709, 1); +lean_inc(x_713); +lean_dec(x_709); +x_659 = x_712; +x_660 = x_713; +goto block_677; } } } } +else +{ +uint8_t x_714; +x_714 = lean_nat_dec_le(x_698, x_698); +if (x_714 == 0) +{ +lean_object* x_715; uint8_t x_716; +lean_dec(x_698); +lean_dec(x_694); +x_715 = lean_array_get_size(x_695); +x_716 = lean_nat_dec_lt(x_699, x_715); +if (x_716 == 0) +{ +lean_object* x_717; +lean_dec(x_715); +lean_dec(x_695); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_717 = lean_box(0); +x_640 = x_717; +x_641 = x_697; +goto block_658; } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); +uint8_t x_718; +x_718 = lean_nat_dec_le(x_715, x_715); +if (x_718 == 0) +{ +lean_object* x_719; +lean_dec(x_715); +lean_dec(x_695); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_719 = lean_box(0); +x_640 = x_719; +x_641 = x_697; +goto block_658; +} +else { -return x_13; +size_t x_720; size_t x_721; lean_object* x_722; lean_object* x_723; +x_720 = 0; +x_721 = lean_usize_of_nat(x_715); +lean_dec(x_715); +x_722 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_723 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__24(x_1, x_2, x_3, x_695, x_720, x_721, x_722, x_5, x_6, x_7, x_697); +lean_dec(x_695); +if (lean_obj_tag(x_723) == 0) +{ +lean_object* x_724; lean_object* x_725; +x_724 = lean_ctor_get(x_723, 0); +lean_inc(x_724); +x_725 = lean_ctor_get(x_723, 1); +lean_inc(x_725); +lean_dec(x_723); +x_640 = x_724; +x_641 = x_725; +goto block_658; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_726; lean_object* x_727; +x_726 = lean_ctor_get(x_723, 0); +lean_inc(x_726); +x_727 = lean_ctor_get(x_723, 1); +lean_inc(x_727); +lean_dec(x_723); +x_659 = x_726; +x_660 = x_727; +goto block_677; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__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) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -lean_inc(x_1); +size_t x_728; size_t x_729; lean_object* x_730; lean_object* x_731; +x_728 = 0; +x_729 = lean_usize_of_nat(x_698); +lean_dec(x_698); +x_730 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_13 = lean_apply_5(x_1, x_10, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -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_13, 1); -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_array_get_size(x_11); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_lt(x_18, x_17); -if (x_19 == 0) +lean_inc(x_2); +x_731 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__25(x_2, x_694, x_728, x_729, x_730, x_5, x_6, x_7, x_697); +lean_dec(x_694); +if (lean_obj_tag(x_731) == 0) { -lean_object* x_20; uint8_t x_21; -lean_dec(x_17); -lean_dec(x_11); -x_20 = lean_array_get_size(x_12); -x_21 = lean_nat_dec_lt(x_18, x_20); -if (x_21 == 0) +lean_object* x_732; lean_object* x_733; uint8_t x_734; +x_732 = lean_ctor_get(x_731, 1); +lean_inc(x_732); +lean_dec(x_731); +x_733 = lean_array_get_size(x_695); +x_734 = lean_nat_dec_lt(x_699, x_733); +if (x_734 == 0) { -lean_object* x_22; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_733); +lean_dec(x_695); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(0); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_640 = x_730; +x_641 = x_732; +goto block_658; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_20, x_20); -if (x_23 == 0) +uint8_t x_735; +x_735 = lean_nat_dec_le(x_733, x_733); +if (x_735 == 0) { -lean_object* x_24; -lean_dec(x_20); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_733); +lean_dec(x_695); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_13, 0, x_24); -return x_13; +x_640 = x_730; +x_641 = x_732; +goto block_658; } else { -size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_13); -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = lean_box(0); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(x_2, x_1, x_3, x_12, x_25, x_26, x_27, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_28; +size_t x_736; lean_object* x_737; +x_736 = lean_usize_of_nat(x_733); +lean_dec(x_733); +lean_inc(x_7); +lean_inc(x_5); +x_737 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__26(x_1, x_2, x_3, x_695, x_728, x_736, x_730, x_5, x_6, x_7, x_732); +lean_dec(x_695); +if (lean_obj_tag(x_737) == 0) +{ +lean_object* x_738; lean_object* x_739; +x_738 = lean_ctor_get(x_737, 0); +lean_inc(x_738); +x_739 = lean_ctor_get(x_737, 1); +lean_inc(x_739); +lean_dec(x_737); +x_640 = x_738; +x_641 = x_739; +goto block_658; +} +else +{ +lean_object* x_740; lean_object* x_741; +x_740 = lean_ctor_get(x_737, 0); +lean_inc(x_740); +x_741 = lean_ctor_get(x_737, 1); +lean_inc(x_741); +lean_dec(x_737); +x_659 = x_740; +x_660 = x_741; +goto block_677; +} } } } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_17, x_17); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -lean_dec(x_17); -lean_dec(x_11); -x_30 = lean_array_get_size(x_12); -x_31 = lean_nat_dec_lt(x_18, x_30); -if (x_31 == 0) -{ -lean_object* x_32; -lean_dec(x_30); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_742; lean_object* x_743; +lean_dec(x_695); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_742 = lean_ctor_get(x_731, 0); +lean_inc(x_742); +x_743 = lean_ctor_get(x_731, 1); +lean_inc(x_743); +lean_dec(x_731); +x_659 = x_742; +x_660 = x_743; +goto block_677; +} +} +} +} +else +{ +lean_object* x_744; lean_object* x_745; +lean_dec(x_695); +lean_dec(x_694); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_744 = lean_ctor_get(x_696, 0); +lean_inc(x_744); +x_745 = lean_ctor_get(x_696, 1); +lean_inc(x_745); +lean_dec(x_696); +x_659 = x_744; +x_660 = x_745; +goto block_677; +} +block_658: +{ +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; uint8_t x_654; +x_642 = lean_st_ref_get(x_7, x_641); +x_643 = lean_ctor_get(x_642, 1); +lean_inc(x_643); +lean_dec(x_642); +x_644 = lean_st_ref_get(x_5, x_643); +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +x_646 = lean_ctor_get(x_644, 1); +lean_inc(x_646); +lean_dec(x_644); +x_647 = lean_ctor_get(x_638, 0); +lean_inc(x_647); +x_648 = lean_ctor_get(x_638, 1); +lean_inc(x_648); +lean_dec(x_638); +x_649 = lean_ctor_get(x_645, 2); +lean_inc(x_649); +lean_dec(x_645); +x_650 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_650, 0, x_647); +lean_ctor_set(x_650, 1, x_648); +lean_ctor_set(x_650, 2, x_649); +x_651 = lean_st_ref_get(x_7, x_646); +lean_dec(x_7); +x_652 = lean_ctor_get(x_651, 1); +lean_inc(x_652); +lean_dec(x_651); +x_653 = lean_st_ref_set(x_5, x_650, x_652); +lean_dec(x_5); +x_654 = !lean_is_exclusive(x_653); +if (x_654 == 0) +{ +lean_object* x_655; +x_655 = lean_ctor_get(x_653, 0); +lean_dec(x_655); +lean_ctor_set(x_653, 0, x_640); +return x_653; } else { -uint8_t x_33; -x_33 = lean_nat_dec_le(x_30, x_30); -if (x_33 == 0) +lean_object* x_656; lean_object* x_657; +x_656 = lean_ctor_get(x_653, 1); +lean_inc(x_656); +lean_dec(x_653); +x_657 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_657, 0, x_640); +lean_ctor_set(x_657, 1, x_656); +return x_657; +} +} +block_677: { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_12); +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; uint8_t x_673; +x_661 = lean_st_ref_get(x_7, x_660); +x_662 = lean_ctor_get(x_661, 1); +lean_inc(x_662); +lean_dec(x_661); +x_663 = lean_st_ref_get(x_5, x_662); +x_664 = lean_ctor_get(x_663, 0); +lean_inc(x_664); +x_665 = lean_ctor_get(x_663, 1); +lean_inc(x_665); +lean_dec(x_663); +x_666 = lean_ctor_get(x_638, 0); +lean_inc(x_666); +x_667 = lean_ctor_get(x_638, 1); +lean_inc(x_667); +lean_dec(x_638); +x_668 = lean_ctor_get(x_664, 2); +lean_inc(x_668); +lean_dec(x_664); +x_669 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_669, 0, x_666); +lean_ctor_set(x_669, 1, x_667); +lean_ctor_set(x_669, 2, x_668); +x_670 = lean_st_ref_get(x_7, x_665); +lean_dec(x_7); +x_671 = lean_ctor_get(x_670, 1); +lean_inc(x_671); +lean_dec(x_670); +x_672 = lean_st_ref_set(x_5, x_669, x_671); +lean_dec(x_5); +x_673 = !lean_is_exclusive(x_672); +if (x_673 == 0) +{ +lean_object* x_674; +x_674 = lean_ctor_get(x_672, 0); +lean_dec(x_674); +lean_ctor_set_tag(x_672, 1); +lean_ctor_set(x_672, 0, x_659); +return x_672; +} +else +{ +lean_object* x_675; lean_object* x_676; +x_675 = lean_ctor_get(x_672, 1); +lean_inc(x_675); +lean_dec(x_672); +x_676 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_676, 0, x_659); +lean_ctor_set(x_676, 1, x_675); +return x_676; +} +} +} +} +else +{ +uint8_t x_746; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -lean_ctor_set(x_13, 0, x_34); -return x_13; +x_746 = !lean_is_exclusive(x_629); +if (x_746 == 0) +{ +return x_629; } else { -size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_13); -x_35 = 0; -x_36 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_37 = lean_box(0); -x_38 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44(x_2, x_1, x_3, x_12, x_35, x_36, x_37, x_5, x_6, x_7, x_15); -lean_dec(x_12); -return x_38; +lean_object* x_747; lean_object* x_748; lean_object* x_749; +x_747 = lean_ctor_get(x_629, 0); +x_748 = lean_ctor_get(x_629, 1); +lean_inc(x_748); +lean_inc(x_747); +lean_dec(x_629); +x_749 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_749, 0, x_747); +lean_ctor_set(x_749, 1, x_748); +return x_749; } } } -else +case 7: { -size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_13); -x_39 = 0; -x_40 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_41 = lean_box(0); +lean_object* x_750; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_750 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_750) == 0) +{ +lean_object* x_751; +x_751 = lean_ctor_get(x_750, 0); +lean_inc(x_751); +if (lean_obj_tag(x_751) == 0) +{ +lean_object* x_752; lean_object* x_753; +lean_dec(x_3); +lean_dec(x_1); +x_752 = lean_ctor_get(x_750, 1); +lean_inc(x_752); +lean_dec(x_750); +x_753 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_752); +return x_753; +} +else +{ +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_780; lean_object* x_781; 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; +x_754 = lean_ctor_get(x_750, 1); +lean_inc(x_754); +lean_dec(x_750); +x_755 = lean_ctor_get(x_751, 0); +lean_inc(x_755); +lean_dec(x_751); +x_756 = lean_st_ref_get(x_7, x_754); +x_757 = lean_ctor_get(x_756, 1); +lean_inc(x_757); +lean_dec(x_756); +x_758 = lean_st_ref_get(x_5, x_757); +x_759 = lean_ctor_get(x_758, 0); +lean_inc(x_759); +x_760 = lean_ctor_get(x_758, 1); +lean_inc(x_760); +lean_dec(x_758); +x_799 = lean_st_ref_get(x_7, x_760); +x_800 = lean_ctor_get(x_799, 1); +lean_inc(x_800); +lean_dec(x_799); +x_801 = lean_st_ref_take(x_5, x_800); +x_802 = lean_ctor_get(x_801, 0); +lean_inc(x_802); +x_803 = lean_ctor_get(x_801, 1); +lean_inc(x_803); +lean_dec(x_801); +x_804 = lean_ctor_get(x_802, 0); +lean_inc(x_804); +x_805 = lean_ctor_get(x_802, 2); +lean_inc(x_805); +lean_dec(x_802); +x_806 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_807 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_807, 0, x_804); +lean_ctor_set(x_807, 1, x_806); +lean_ctor_set(x_807, 2, x_805); +x_808 = lean_st_ref_set(x_5, x_807, x_803); +x_809 = lean_ctor_get(x_808, 1); +lean_inc(x_809); +lean_dec(x_808); +x_810 = l_Lean_Compiler_visitMatch(x_9, x_755, x_5, x_6, x_7, x_809); +x_811 = lean_ctor_get(x_810, 0); +lean_inc(x_811); +x_812 = lean_ctor_get(x_811, 1); +lean_inc(x_812); +x_813 = lean_ctor_get(x_810, 1); +lean_inc(x_813); +lean_dec(x_810); +x_814 = lean_ctor_get(x_811, 0); +lean_inc(x_814); +lean_dec(x_811); +x_815 = lean_ctor_get(x_812, 0); +lean_inc(x_815); +x_816 = lean_ctor_get(x_812, 1); +lean_inc(x_816); +lean_dec(x_812); +lean_inc(x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(x_1, x_11, x_39, x_40, x_41, x_5, x_6, x_7, x_15); -lean_dec(x_11); -if (lean_obj_tag(x_42) == 0) +x_817 = lean_apply_5(x_2, x_814, x_5, x_6, x_7, x_813); +if (lean_obj_tag(x_817) == 0) { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_object* x_818; lean_object* x_819; lean_object* x_820; uint8_t x_821; +x_818 = lean_ctor_get(x_817, 1); +lean_inc(x_818); +lean_dec(x_817); +x_819 = lean_array_get_size(x_815); +x_820 = lean_unsigned_to_nat(0u); +x_821 = lean_nat_dec_lt(x_820, x_819); +if (x_821 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_42, 1); -x_45 = lean_ctor_get(x_42, 0); -lean_dec(x_45); -x_46 = lean_array_get_size(x_12); -x_47 = lean_nat_dec_lt(x_18, x_46); -if (x_47 == 0) +lean_object* x_822; uint8_t x_823; +lean_dec(x_819); +lean_dec(x_815); +x_822 = lean_array_get_size(x_816); +x_823 = lean_nat_dec_lt(x_820, x_822); +if (x_823 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_824; +lean_dec(x_822); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_824 = lean_box(0); +x_761 = x_824; +x_762 = x_818; +goto block_779; } else { -uint8_t x_48; -x_48 = lean_nat_dec_le(x_46, x_46); -if (x_48 == 0) +uint8_t x_825; +x_825 = lean_nat_dec_le(x_822, x_822); +if (x_825 == 0) { -lean_dec(x_46); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_826; +lean_dec(x_822); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -lean_ctor_set(x_42, 0, x_41); -return x_42; +x_826 = lean_box(0); +x_761 = x_826; +x_762 = x_818; +goto block_779; } else { -size_t x_49; lean_object* x_50; -lean_free_object(x_42); -x_49 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(x_2, x_1, x_3, x_12, x_39, x_49, x_41, x_5, x_6, x_7, x_44); -lean_dec(x_12); -return x_50; +size_t x_827; size_t x_828; lean_object* x_829; lean_object* x_830; +x_827 = 0; +x_828 = lean_usize_of_nat(x_822); +lean_dec(x_822); +x_829 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_830 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__27(x_1, x_2, x_3, x_816, x_827, x_828, x_829, x_5, x_6, x_7, x_818); +lean_dec(x_816); +if (lean_obj_tag(x_830) == 0) +{ +lean_object* x_831; lean_object* x_832; +x_831 = lean_ctor_get(x_830, 0); +lean_inc(x_831); +x_832 = lean_ctor_get(x_830, 1); +lean_inc(x_832); +lean_dec(x_830); +x_761 = x_831; +x_762 = x_832; +goto block_779; +} +else +{ +lean_object* x_833; lean_object* x_834; +x_833 = lean_ctor_get(x_830, 0); +lean_inc(x_833); +x_834 = lean_ctor_get(x_830, 1); +lean_inc(x_834); +lean_dec(x_830); +x_780 = x_833; +x_781 = x_834; +goto block_798; +} } } } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); -lean_dec(x_42); -x_52 = lean_array_get_size(x_12); -x_53 = lean_nat_dec_lt(x_18, x_52); -if (x_53 == 0) +uint8_t x_835; +x_835 = lean_nat_dec_le(x_819, x_819); +if (x_835 == 0) { -lean_object* x_54; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_836; uint8_t x_837; +lean_dec(x_819); +lean_dec(x_815); +x_836 = lean_array_get_size(x_816); +x_837 = lean_nat_dec_lt(x_820, x_836); +if (x_837 == 0) +{ +lean_object* x_838; +lean_dec(x_836); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_51); -return x_54; +x_838 = lean_box(0); +x_761 = x_838; +x_762 = x_818; +goto block_779; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_52, x_52); -if (x_55 == 0) +uint8_t x_839; +x_839 = lean_nat_dec_le(x_836, x_836); +if (x_839 == 0) { -lean_object* x_56; -lean_dec(x_52); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_840; +lean_dec(x_836); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -else -{ -size_t x_57; lean_object* x_58; -x_57 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(x_2, x_1, x_3, x_12, x_39, x_57, x_41, x_5, x_6, x_7, x_51); -lean_dec(x_12); -return x_58; -} -} -} +x_840 = lean_box(0); +x_761 = x_840; +x_762 = x_818; +goto block_779; } else { -uint8_t x_59; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_42); -if (x_59 == 0) +size_t x_841; size_t x_842; lean_object* x_843; lean_object* x_844; +x_841 = 0; +x_842 = lean_usize_of_nat(x_836); +lean_dec(x_836); +x_843 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_844 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__28(x_1, x_2, x_3, x_816, x_841, x_842, x_843, x_5, x_6, x_7, x_818); +lean_dec(x_816); +if (lean_obj_tag(x_844) == 0) { -return x_42; +lean_object* x_845; lean_object* x_846; +x_845 = lean_ctor_get(x_844, 0); +lean_inc(x_845); +x_846 = lean_ctor_get(x_844, 1); +lean_inc(x_846); +lean_dec(x_844); +x_761 = x_845; +x_762 = x_846; +goto block_779; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_42, 0); -x_61 = lean_ctor_get(x_42, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_42); -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_847; lean_object* x_848; +x_847 = lean_ctor_get(x_844, 0); +lean_inc(x_847); +x_848 = lean_ctor_get(x_844, 1); +lean_inc(x_848); +lean_dec(x_844); +x_780 = x_847; +x_781 = x_848; +goto block_798; } } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_13, 1); -lean_inc(x_63); -lean_dec(x_13); -x_64 = lean_array_get_size(x_11); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_lt(x_65, x_64); -if (x_66 == 0) +size_t x_849; size_t x_850; lean_object* x_851; lean_object* x_852; +x_849 = 0; +x_850 = lean_usize_of_nat(x_819); +lean_dec(x_819); +x_851 = lean_box(0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_852 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__29(x_2, x_815, x_849, x_850, x_851, x_5, x_6, x_7, x_818); +lean_dec(x_815); +if (lean_obj_tag(x_852) == 0) { -lean_object* x_67; uint8_t x_68; -lean_dec(x_64); -lean_dec(x_11); -x_67 = lean_array_get_size(x_12); -x_68 = lean_nat_dec_lt(x_65, x_67); -if (x_68 == 0) +lean_object* x_853; lean_object* x_854; uint8_t x_855; +x_853 = lean_ctor_get(x_852, 1); +lean_inc(x_853); +lean_dec(x_852); +x_854 = lean_array_get_size(x_816); +x_855 = lean_nat_dec_lt(x_820, x_854); +if (x_855 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_854); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = lean_box(0); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_63); -return x_70; +x_761 = x_851; +x_762 = x_853; +goto block_779; } else { -uint8_t x_71; -x_71 = lean_nat_dec_le(x_67, x_67); -if (x_71 == 0) +uint8_t x_856; +x_856 = lean_nat_dec_le(x_854, x_854); +if (x_856 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_67); -lean_dec(x_12); -lean_dec(x_7); +lean_dec(x_854); +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_63); -return x_73; +x_761 = x_851; +x_762 = x_853; +goto block_779; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -x_74 = 0; -x_75 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(x_2, x_1, x_3, x_12, x_74, x_75, x_76, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_77; +size_t x_857; lean_object* x_858; +x_857 = lean_usize_of_nat(x_854); +lean_dec(x_854); +lean_inc(x_7); +lean_inc(x_5); +x_858 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__30(x_1, x_2, x_3, x_816, x_849, x_857, x_851, x_5, x_6, x_7, x_853); +lean_dec(x_816); +if (lean_obj_tag(x_858) == 0) +{ +lean_object* x_859; lean_object* x_860; +x_859 = lean_ctor_get(x_858, 0); +lean_inc(x_859); +x_860 = lean_ctor_get(x_858, 1); +lean_inc(x_860); +lean_dec(x_858); +x_761 = x_859; +x_762 = x_860; +goto block_779; +} +else +{ +lean_object* x_861; lean_object* x_862; +x_861 = lean_ctor_get(x_858, 0); +lean_inc(x_861); +x_862 = lean_ctor_get(x_858, 1); +lean_inc(x_862); +lean_dec(x_858); +x_780 = x_861; +x_781 = x_862; +goto block_798; +} } } } else { -uint8_t x_78; -x_78 = lean_nat_dec_le(x_64, x_64); -if (x_78 == 0) -{ -lean_object* x_79; uint8_t x_80; -lean_dec(x_64); -lean_dec(x_11); -x_79 = lean_array_get_size(x_12); -x_80 = lean_nat_dec_lt(x_65, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_79); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_863; lean_object* x_864; +lean_dec(x_816); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_63); -return x_82; +x_863 = lean_ctor_get(x_852, 0); +lean_inc(x_863); +x_864 = lean_ctor_get(x_852, 1); +lean_inc(x_864); +lean_dec(x_852); +x_780 = x_863; +x_781 = x_864; +goto block_798; +} +} +} +} +else +{ +lean_object* x_865; lean_object* x_866; +lean_dec(x_816); +lean_dec(x_815); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_865 = lean_ctor_get(x_817, 0); +lean_inc(x_865); +x_866 = lean_ctor_get(x_817, 1); +lean_inc(x_866); +lean_dec(x_817); +x_780 = x_865; +x_781 = x_866; +goto block_798; +} +block_779: +{ +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; +x_763 = lean_st_ref_get(x_7, x_762); +x_764 = lean_ctor_get(x_763, 1); +lean_inc(x_764); +lean_dec(x_763); +x_765 = lean_st_ref_get(x_5, 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); +lean_dec(x_765); +x_768 = lean_ctor_get(x_759, 0); +lean_inc(x_768); +x_769 = lean_ctor_get(x_759, 1); +lean_inc(x_769); +lean_dec(x_759); +x_770 = lean_ctor_get(x_766, 2); +lean_inc(x_770); +lean_dec(x_766); +x_771 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_771, 0, x_768); +lean_ctor_set(x_771, 1, x_769); +lean_ctor_set(x_771, 2, x_770); +x_772 = lean_st_ref_get(x_7, x_767); +lean_dec(x_7); +x_773 = lean_ctor_get(x_772, 1); +lean_inc(x_773); +lean_dec(x_772); +x_774 = lean_st_ref_set(x_5, x_771, x_773); +lean_dec(x_5); +x_775 = !lean_is_exclusive(x_774); +if (x_775 == 0) +{ +lean_object* x_776; +x_776 = lean_ctor_get(x_774, 0); +lean_dec(x_776); +lean_ctor_set(x_774, 0, x_761); +return x_774; } else { -uint8_t x_83; -x_83 = lean_nat_dec_le(x_79, x_79); -if (x_83 == 0) +lean_object* x_777; lean_object* x_778; +x_777 = lean_ctor_get(x_774, 1); +lean_inc(x_777); +lean_dec(x_774); +x_778 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_778, 0, x_761); +lean_ctor_set(x_778, 1, x_777); +return x_778; +} +} +block_798: { -lean_object* x_84; lean_object* x_85; -lean_dec(x_79); -lean_dec(x_12); +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; uint8_t x_794; +x_782 = lean_st_ref_get(x_7, x_781); +x_783 = lean_ctor_get(x_782, 1); +lean_inc(x_783); +lean_dec(x_782); +x_784 = lean_st_ref_get(x_5, x_783); +x_785 = lean_ctor_get(x_784, 0); +lean_inc(x_785); +x_786 = lean_ctor_get(x_784, 1); +lean_inc(x_786); +lean_dec(x_784); +x_787 = lean_ctor_get(x_759, 0); +lean_inc(x_787); +x_788 = lean_ctor_get(x_759, 1); +lean_inc(x_788); +lean_dec(x_759); +x_789 = lean_ctor_get(x_785, 2); +lean_inc(x_789); +lean_dec(x_785); +x_790 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_790, 0, x_787); +lean_ctor_set(x_790, 1, x_788); +lean_ctor_set(x_790, 2, x_789); +x_791 = lean_st_ref_get(x_7, x_786); +lean_dec(x_7); +x_792 = lean_ctor_get(x_791, 1); +lean_inc(x_792); +lean_dec(x_791); +x_793 = lean_st_ref_set(x_5, x_790, x_792); +lean_dec(x_5); +x_794 = !lean_is_exclusive(x_793); +if (x_794 == 0) +{ +lean_object* x_795; +x_795 = lean_ctor_get(x_793, 0); +lean_dec(x_795); +lean_ctor_set_tag(x_793, 1); +lean_ctor_set(x_793, 0, x_780); +return x_793; +} +else +{ +lean_object* x_796; lean_object* x_797; +x_796 = lean_ctor_get(x_793, 1); +lean_inc(x_796); +lean_dec(x_793); +x_797 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_797, 0, x_780); +lean_ctor_set(x_797, 1, x_796); +return x_797; +} +} +} +} +else +{ +uint8_t x_867; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_63); -return x_85; +x_867 = !lean_is_exclusive(x_750); +if (x_867 == 0) +{ +return x_750; } else { -size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = 0; -x_87 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_88 = lean_box(0); -x_89 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44(x_2, x_1, x_3, x_12, x_86, x_87, x_88, x_5, x_6, x_7, x_63); -lean_dec(x_12); -return x_89; +lean_object* x_868; lean_object* x_869; lean_object* x_870; +x_868 = lean_ctor_get(x_750, 0); +x_869 = lean_ctor_get(x_750, 1); +lean_inc(x_869); +lean_inc(x_868); +lean_dec(x_750); +x_870 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_870, 0, x_868); +lean_ctor_set(x_870, 1, x_869); +return x_870; } } } -else +case 8: { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; -x_90 = 0; -x_91 = lean_usize_of_nat(x_64); -lean_dec(x_64); -x_92 = lean_box(0); +lean_object* x_871; +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_9); +x_871 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_871) == 0) +{ +lean_object* x_872; +x_872 = lean_ctor_get(x_871, 0); +lean_inc(x_872); +if (lean_obj_tag(x_872) == 0) +{ +lean_object* x_873; lean_object* x_874; +lean_dec(x_3); +lean_dec(x_1); +x_873 = lean_ctor_get(x_871, 1); +lean_inc(x_873); +lean_dec(x_871); +x_874 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_873); +return x_874; +} +else +{ +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_901; lean_object* x_902; 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; +x_875 = lean_ctor_get(x_871, 1); +lean_inc(x_875); +lean_dec(x_871); +x_876 = lean_ctor_get(x_872, 0); +lean_inc(x_876); +lean_dec(x_872); +x_877 = lean_st_ref_get(x_7, x_875); +x_878 = lean_ctor_get(x_877, 1); +lean_inc(x_878); +lean_dec(x_877); +x_879 = lean_st_ref_get(x_5, x_878); +x_880 = lean_ctor_get(x_879, 0); +lean_inc(x_880); +x_881 = lean_ctor_get(x_879, 1); +lean_inc(x_881); +lean_dec(x_879); +x_920 = lean_st_ref_get(x_7, x_881); +x_921 = lean_ctor_get(x_920, 1); +lean_inc(x_921); +lean_dec(x_920); +x_922 = lean_st_ref_take(x_5, x_921); +x_923 = lean_ctor_get(x_922, 0); +lean_inc(x_923); +x_924 = lean_ctor_get(x_922, 1); +lean_inc(x_924); +lean_dec(x_922); +x_925 = lean_ctor_get(x_923, 0); +lean_inc(x_925); +x_926 = lean_ctor_get(x_923, 2); +lean_inc(x_926); +lean_dec(x_923); +x_927 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_928 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_928, 0, x_925); +lean_ctor_set(x_928, 1, x_927); +lean_ctor_set(x_928, 2, x_926); +x_929 = lean_st_ref_set(x_5, x_928, x_924); +x_930 = lean_ctor_get(x_929, 1); +lean_inc(x_930); +lean_dec(x_929); +x_931 = l_Lean_Compiler_visitMatch(x_9, x_876, x_5, x_6, x_7, x_930); +x_932 = lean_ctor_get(x_931, 0); +lean_inc(x_932); +x_933 = lean_ctor_get(x_932, 1); +lean_inc(x_933); +x_934 = lean_ctor_get(x_931, 1); +lean_inc(x_934); +lean_dec(x_931); +x_935 = lean_ctor_get(x_932, 0); +lean_inc(x_935); +lean_dec(x_932); +x_936 = lean_ctor_get(x_933, 0); +lean_inc(x_936); +x_937 = lean_ctor_get(x_933, 1); +lean_inc(x_937); +lean_dec(x_933); +lean_inc(x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); -x_93 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(x_1, x_11, x_90, x_91, x_92, x_5, x_6, x_7, x_63); -lean_dec(x_11); -if (lean_obj_tag(x_93) == 0) +x_938 = lean_apply_5(x_2, x_935, x_5, x_6, x_7, x_934); +if (lean_obj_tag(x_938) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; -} else { - lean_dec_ref(x_93); - x_95 = lean_box(0); -} -x_96 = lean_array_get_size(x_12); -x_97 = lean_nat_dec_lt(x_65, x_96); -if (x_97 == 0) +lean_object* x_939; lean_object* x_940; lean_object* x_941; uint8_t x_942; +x_939 = lean_ctor_get(x_938, 1); +lean_inc(x_939); +lean_dec(x_938); +x_940 = lean_array_get_size(x_936); +x_941 = lean_unsigned_to_nat(0u); +x_942 = lean_nat_dec_lt(x_941, x_940); +if (x_942 == 0) { -lean_object* x_98; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_943; uint8_t x_944; +lean_dec(x_940); +lean_dec(x_936); +x_943 = lean_array_get_size(x_937); +x_944 = lean_nat_dec_lt(x_941, x_943); +if (x_944 == 0) +{ +lean_object* x_945; +lean_dec(x_943); +lean_dec(x_937); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_95; -} -lean_ctor_set(x_98, 0, x_92); -lean_ctor_set(x_98, 1, x_94); -return x_98; +x_945 = lean_box(0); +x_882 = x_945; +x_883 = x_939; +goto block_900; } else { -uint8_t x_99; -x_99 = lean_nat_dec_le(x_96, x_96); -if (x_99 == 0) -{ -lean_object* x_100; -lean_dec(x_96); -lean_dec(x_12); -lean_dec(x_7); +uint8_t x_946; +x_946 = lean_nat_dec_le(x_943, x_943); +if (x_946 == 0) +{ +lean_object* x_947; +lean_dec(x_943); +lean_dec(x_937); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_95)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_95; +x_947 = lean_box(0); +x_882 = x_947; +x_883 = x_939; +goto block_900; } -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_94); -return x_100; +else +{ +size_t x_948; size_t x_949; lean_object* x_950; lean_object* x_951; +x_948 = 0; +x_949 = lean_usize_of_nat(x_943); +lean_dec(x_943); +x_950 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_951 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__31(x_1, x_2, x_3, x_937, x_948, x_949, x_950, x_5, x_6, x_7, x_939); +lean_dec(x_937); +if (lean_obj_tag(x_951) == 0) +{ +lean_object* x_952; lean_object* x_953; +x_952 = lean_ctor_get(x_951, 0); +lean_inc(x_952); +x_953 = lean_ctor_get(x_951, 1); +lean_inc(x_953); +lean_dec(x_951); +x_882 = x_952; +x_883 = x_953; +goto block_900; } else { -size_t x_101; lean_object* x_102; -lean_dec(x_95); -x_101 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(x_2, x_1, x_3, x_12, x_90, x_101, x_92, x_5, x_6, x_7, x_94); -lean_dec(x_12); -return x_102; +lean_object* x_954; lean_object* x_955; +x_954 = lean_ctor_get(x_951, 0); +lean_inc(x_954); +x_955 = lean_ctor_get(x_951, 1); +lean_inc(x_955); +lean_dec(x_951); +x_901 = x_954; +x_902 = x_955; +goto block_919; +} } } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_12); -lean_dec(x_7); +uint8_t x_956; +x_956 = lean_nat_dec_le(x_940, x_940); +if (x_956 == 0) +{ +lean_object* x_957; uint8_t x_958; +lean_dec(x_940); +lean_dec(x_936); +x_957 = lean_array_get_size(x_937); +x_958 = lean_nat_dec_lt(x_941, x_957); +if (x_958 == 0) +{ +lean_object* x_959; +lean_dec(x_957); +lean_dec(x_937); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_103 = lean_ctor_get(x_93, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_93, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_105 = x_93; -} else { - lean_dec_ref(x_93); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; -} -} -} -} +x_959 = lean_box(0); +x_882 = x_959; +x_883 = x_939; +goto block_900; } else { -uint8_t x_107; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); +uint8_t x_960; +x_960 = lean_nat_dec_le(x_957, x_957); +if (x_960 == 0) +{ +lean_object* x_961; +lean_dec(x_957); +lean_dec(x_937); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_107 = !lean_is_exclusive(x_13); -if (x_107 == 0) +x_961 = lean_box(0); +x_882 = x_961; +x_883 = x_939; +goto block_900; +} +else { -return x_13; +size_t x_962; size_t x_963; lean_object* x_964; lean_object* x_965; +x_962 = 0; +x_963 = lean_usize_of_nat(x_957); +lean_dec(x_957); +x_964 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_965 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__32(x_1, x_2, x_3, x_937, x_962, x_963, x_964, x_5, x_6, x_7, x_939); +lean_dec(x_937); +if (lean_obj_tag(x_965) == 0) +{ +lean_object* x_966; lean_object* x_967; +x_966 = lean_ctor_get(x_965, 0); +lean_inc(x_966); +x_967 = lean_ctor_get(x_965, 1); +lean_inc(x_967); +lean_dec(x_965); +x_882 = x_966; +x_883 = x_967; +goto block_900; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_13, 0); -x_109 = lean_ctor_get(x_13, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_13); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_968; lean_object* x_969; +x_968 = lean_ctor_get(x_965, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_965, 1); +lean_inc(x_969); +lean_dec(x_965); +x_901 = x_968; +x_902 = x_969; +goto block_919; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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: +else { -lean_object* x_9; -x_9 = l_Lean_Expr_consumeMData(x_4); -switch (lean_obj_tag(x_9)) { -case 0: +size_t x_970; size_t x_971; lean_object* x_972; lean_object* x_973; +x_970 = 0; +x_971 = lean_usize_of_nat(x_940); +lean_dec(x_940); +x_972 = lean_box(0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_973 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__33(x_2, x_936, x_970, x_971, x_972, x_5, x_6, x_7, x_939); +lean_dec(x_936); +if (lean_obj_tag(x_973) == 0) { -lean_object* x_10; lean_object* x_11; -lean_dec(x_9); +lean_object* x_974; lean_object* x_975; uint8_t x_976; +x_974 = lean_ctor_get(x_973, 1); +lean_inc(x_974); +lean_dec(x_973); +x_975 = lean_array_get_size(x_937); +x_976 = lean_nat_dec_lt(x_941, x_975); +if (x_976 == 0) +{ +lean_dec(x_975); +lean_dec(x_937); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_11 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_10, x_5, x_6, x_7, x_8); -return x_11; +x_882 = x_972; +x_883 = x_974; +goto block_900; } -case 2: +else { -lean_object* x_12; lean_object* x_13; -lean_dec(x_9); +uint8_t x_977; +x_977 = lean_nat_dec_le(x_975, x_975); +if (x_977 == 0) +{ +lean_dec(x_975); +lean_dec(x_937); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_12 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_13 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_12, x_5, x_6, x_7, x_8); -return x_13; +x_882 = x_972; +x_883 = x_974; +goto block_900; } -case 5: -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -switch (lean_obj_tag(x_14)) { -case 0: +else { -lean_object* x_15; -lean_dec(x_14); +size_t x_978; lean_object* x_979; +x_978 = lean_usize_of_nat(x_975); +lean_dec(x_975); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_15 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_15) == 0) +lean_inc(x_5); +x_979 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__34(x_1, x_2, x_3, x_937, x_970, x_978, x_972, x_5, x_6, x_7, x_974); +lean_dec(x_937); +if (lean_obj_tag(x_979) == 0) { -lean_object* x_16; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -if (lean_obj_tag(x_16) == 0) +lean_object* x_980; lean_object* x_981; +x_980 = lean_ctor_get(x_979, 0); +lean_inc(x_980); +x_981 = lean_ctor_get(x_979, 1); +lean_inc(x_981); +lean_dec(x_979); +x_882 = x_980; +x_883 = x_981; +goto block_900; +} +else { -lean_object* x_17; lean_object* x_18; +lean_object* x_982; lean_object* x_983; +x_982 = lean_ctor_get(x_979, 0); +lean_inc(x_982); +x_983 = lean_ctor_get(x_979, 1); +lean_inc(x_983); +lean_dec(x_979); +x_901 = x_982; +x_902 = x_983; +goto block_919; +} +} +} +} +else +{ +lean_object* x_984; lean_object* x_985; +lean_dec(x_937); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_17); -return x_18; +x_984 = lean_ctor_get(x_973, 0); +lean_inc(x_984); +x_985 = lean_ctor_get(x_973, 1); +lean_inc(x_985); +lean_dec(x_973); +x_901 = x_984; +x_902 = x_985; +goto block_919; +} +} +} +} +else +{ +lean_object* x_986; lean_object* x_987; +lean_dec(x_937); +lean_dec(x_936); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_986 = lean_ctor_get(x_938, 0); +lean_inc(x_986); +x_987 = lean_ctor_get(x_938, 1); +lean_inc(x_987); +lean_dec(x_938); +x_901 = x_986; +x_902 = x_987; +goto block_919; +} +block_900: +{ +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; uint8_t x_896; +x_884 = lean_st_ref_get(x_7, x_883); +x_885 = lean_ctor_get(x_884, 1); +lean_inc(x_885); +lean_dec(x_884); +x_886 = lean_st_ref_get(x_5, x_885); +x_887 = lean_ctor_get(x_886, 0); +lean_inc(x_887); +x_888 = lean_ctor_get(x_886, 1); +lean_inc(x_888); +lean_dec(x_886); +x_889 = lean_ctor_get(x_880, 0); +lean_inc(x_889); +x_890 = lean_ctor_get(x_880, 1); +lean_inc(x_890); +lean_dec(x_880); +x_891 = lean_ctor_get(x_887, 2); +lean_inc(x_891); +lean_dec(x_887); +x_892 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_892, 0, x_889); +lean_ctor_set(x_892, 1, x_890); +lean_ctor_set(x_892, 2, x_891); +x_893 = lean_st_ref_get(x_7, x_888); +lean_dec(x_7); +x_894 = lean_ctor_get(x_893, 1); +lean_inc(x_894); +lean_dec(x_893); +x_895 = lean_st_ref_set(x_5, x_892, x_894); +lean_dec(x_5); +x_896 = !lean_is_exclusive(x_895); +if (x_896 == 0) +{ +lean_object* x_897; +x_897 = lean_ctor_get(x_895, 0); +lean_dec(x_897); +lean_ctor_set(x_895, 0, x_882); +return x_895; } 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_ctor_get(x_15, 1); -lean_inc(x_19); -lean_dec(x_15); -x_20 = lean_ctor_get(x_16, 0); -lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_21, 0, x_9); -lean_closure_set(x_21, 1, x_20); -x_22 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__1), 8, 3); -lean_closure_set(x_22, 0, x_2); -lean_closure_set(x_22, 1, x_1); -lean_closure_set(x_22, 2, x_3); -x_23 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_23, 0, x_21); -lean_closure_set(x_23, 1, x_22); -x_24 = l_Lean_Compiler_withNewScopeImp___rarg(x_23, x_5, x_6, x_7, x_19); -return x_24; +lean_object* x_898; lean_object* x_899; +x_898 = lean_ctor_get(x_895, 1); +lean_inc(x_898); +lean_dec(x_895); +x_899 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_899, 0, x_882); +lean_ctor_set(x_899, 1, x_898); +return x_899; } } -else +block_919: { -uint8_t x_25; +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; uint8_t x_915; +x_903 = lean_st_ref_get(x_7, x_902); +x_904 = lean_ctor_get(x_903, 1); +lean_inc(x_904); +lean_dec(x_903); +x_905 = lean_st_ref_get(x_5, x_904); +x_906 = lean_ctor_get(x_905, 0); +lean_inc(x_906); +x_907 = lean_ctor_get(x_905, 1); +lean_inc(x_907); +lean_dec(x_905); +x_908 = lean_ctor_get(x_880, 0); +lean_inc(x_908); +x_909 = lean_ctor_get(x_880, 1); +lean_inc(x_909); +lean_dec(x_880); +x_910 = lean_ctor_get(x_906, 2); +lean_inc(x_910); +lean_dec(x_906); +x_911 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_911, 0, x_908); +lean_ctor_set(x_911, 1, x_909); +lean_ctor_set(x_911, 2, x_910); +x_912 = lean_st_ref_get(x_7, x_907); +lean_dec(x_7); +x_913 = lean_ctor_get(x_912, 1); +lean_inc(x_913); +lean_dec(x_912); +x_914 = lean_st_ref_set(x_5, x_911, x_913); +lean_dec(x_5); +x_915 = !lean_is_exclusive(x_914); +if (x_915 == 0) +{ +lean_object* x_916; +x_916 = lean_ctor_get(x_914, 0); +lean_dec(x_916); +lean_ctor_set_tag(x_914, 1); +lean_ctor_set(x_914, 0, x_901); +return x_914; +} +else +{ +lean_object* x_917; lean_object* x_918; +x_917 = lean_ctor_get(x_914, 1); +lean_inc(x_917); +lean_dec(x_914); +x_918 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_918, 0, x_901); +lean_ctor_set(x_918, 1, x_917); +return x_918; +} +} +} +} +else +{ +uint8_t x_988; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -30626,280 +29598,516 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +x_988 = !lean_is_exclusive(x_871); +if (x_988 == 0) { -return x_15; +return x_871; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -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_object* x_989; lean_object* x_990; lean_object* x_991; +x_989 = lean_ctor_get(x_871, 0); +x_990 = lean_ctor_get(x_871, 1); +lean_inc(x_990); +lean_inc(x_989); +lean_dec(x_871); +x_991 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_991, 0, x_989); +lean_ctor_set(x_991, 1, x_990); +return x_991; } } } -case 1: +case 9: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_3); -x_29 = lean_ctor_get(x_9, 1); -lean_inc(x_29); -x_30 = lean_ctor_get(x_14, 0); -lean_inc(x_30); +lean_object* x_992; lean_dec(x_14); lean_inc(x_7); lean_inc(x_6); +lean_inc(x_9); +x_992 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_992) == 0) +{ +lean_object* x_993; +x_993 = lean_ctor_get(x_992, 0); +lean_inc(x_993); +if (lean_obj_tag(x_993) == 0) +{ +lean_object* x_994; lean_object* x_995; +lean_dec(x_3); +lean_dec(x_1); +x_994 = lean_ctor_get(x_992, 1); +lean_inc(x_994); +lean_dec(x_992); +x_995 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_994); +return x_995; +} +else +{ +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_1022; lean_object* x_1023; 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; +x_996 = lean_ctor_get(x_992, 1); +lean_inc(x_996); +lean_dec(x_992); +x_997 = lean_ctor_get(x_993, 0); +lean_inc(x_997); +lean_dec(x_993); +x_998 = lean_st_ref_get(x_7, x_996); +x_999 = lean_ctor_get(x_998, 1); +lean_inc(x_999); +lean_dec(x_998); +x_1000 = lean_st_ref_get(x_5, x_999); +x_1001 = lean_ctor_get(x_1000, 0); +lean_inc(x_1001); +x_1002 = lean_ctor_get(x_1000, 1); +lean_inc(x_1002); +lean_dec(x_1000); +x_1041 = lean_st_ref_get(x_7, x_1002); +x_1042 = lean_ctor_get(x_1041, 1); +lean_inc(x_1042); +lean_dec(x_1041); +x_1043 = lean_st_ref_take(x_5, 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); +lean_dec(x_1043); +x_1046 = lean_ctor_get(x_1044, 0); +lean_inc(x_1046); +x_1047 = lean_ctor_get(x_1044, 2); +lean_inc(x_1047); +lean_dec(x_1044); +x_1048 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1049 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1049, 0, x_1046); +lean_ctor_set(x_1049, 1, x_1048); +lean_ctor_set(x_1049, 2, x_1047); +x_1050 = lean_st_ref_set(x_5, x_1049, x_1045); +x_1051 = lean_ctor_get(x_1050, 1); +lean_inc(x_1051); +lean_dec(x_1050); +x_1052 = l_Lean_Compiler_visitMatch(x_9, x_997, x_5, x_6, x_7, x_1051); +x_1053 = lean_ctor_get(x_1052, 0); +lean_inc(x_1053); +x_1054 = lean_ctor_get(x_1053, 1); +lean_inc(x_1054); +x_1055 = lean_ctor_get(x_1052, 1); +lean_inc(x_1055); +lean_dec(x_1052); +x_1056 = lean_ctor_get(x_1053, 0); +lean_inc(x_1056); +lean_dec(x_1053); +x_1057 = lean_ctor_get(x_1054, 0); +lean_inc(x_1057); +x_1058 = lean_ctor_get(x_1054, 1); +lean_inc(x_1058); +lean_dec(x_1054); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_31 = lean_apply_6(x_1, x_30, x_9, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_31) == 0) +x_1059 = lean_apply_5(x_2, x_1056, x_5, x_6, x_7, x_1055); +if (lean_obj_tag(x_1059) == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_apply_5(x_2, x_29, x_5, x_6, x_7, x_32); -return x_33; +lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; uint8_t x_1063; +x_1060 = lean_ctor_get(x_1059, 1); +lean_inc(x_1060); +lean_dec(x_1059); +x_1061 = lean_array_get_size(x_1057); +x_1062 = lean_unsigned_to_nat(0u); +x_1063 = lean_nat_dec_lt(x_1062, x_1061); +if (x_1063 == 0) +{ +lean_object* x_1064; uint8_t x_1065; +lean_dec(x_1061); +lean_dec(x_1057); +x_1064 = lean_array_get_size(x_1058); +x_1065 = lean_nat_dec_lt(x_1062, x_1064); +if (x_1065 == 0) +{ +lean_object* x_1066; +lean_dec(x_1064); +lean_dec(x_1058); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1066 = lean_box(0); +x_1003 = x_1066; +x_1004 = x_1060; +goto block_1021; } else { -uint8_t x_34; -lean_dec(x_29); -lean_dec(x_7); +uint8_t x_1067; +x_1067 = lean_nat_dec_le(x_1064, x_1064); +if (x_1067 == 0) +{ +lean_object* x_1068; +lean_dec(x_1064); +lean_dec(x_1058); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -x_34 = !lean_is_exclusive(x_31); -if (x_34 == 0) +lean_dec(x_1); +x_1068 = lean_box(0); +x_1003 = x_1068; +x_1004 = x_1060; +goto block_1021; +} +else { -return x_31; +size_t x_1069; size_t x_1070; lean_object* x_1071; lean_object* x_1072; +x_1069 = 0; +x_1070 = lean_usize_of_nat(x_1064); +lean_dec(x_1064); +x_1071 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1072 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__35(x_1, x_2, x_3, x_1058, x_1069, x_1070, x_1071, x_5, x_6, x_7, x_1060); +lean_dec(x_1058); +if (lean_obj_tag(x_1072) == 0) +{ +lean_object* x_1073; lean_object* x_1074; +x_1073 = lean_ctor_get(x_1072, 0); +lean_inc(x_1073); +x_1074 = lean_ctor_get(x_1072, 1); +lean_inc(x_1074); +lean_dec(x_1072); +x_1003 = x_1073; +x_1004 = x_1074; +goto block_1021; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_31, 0); -x_36 = lean_ctor_get(x_31, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_31); -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_object* x_1075; lean_object* x_1076; +x_1075 = lean_ctor_get(x_1072, 0); +lean_inc(x_1075); +x_1076 = lean_ctor_get(x_1072, 1); +lean_inc(x_1076); +lean_dec(x_1072); +x_1022 = x_1075; +x_1023 = x_1076; +goto block_1040; } } } -case 2: +} +else { -lean_object* x_38; -lean_dec(x_14); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_38 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_38) == 0) +uint8_t x_1077; +x_1077 = lean_nat_dec_le(x_1061, x_1061); +if (x_1077 == 0) { -lean_object* x_39; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -if (lean_obj_tag(x_39) == 0) +lean_object* x_1078; uint8_t x_1079; +lean_dec(x_1061); +lean_dec(x_1057); +x_1078 = lean_array_get_size(x_1058); +x_1079 = lean_nat_dec_lt(x_1062, x_1078); +if (x_1079 == 0) { -lean_object* x_40; lean_object* x_41; +lean_object* x_1080; +lean_dec(x_1078); +lean_dec(x_1058); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_40); -return x_41; +x_1080 = lean_box(0); +x_1003 = x_1080; +x_1004 = x_1060; +goto block_1021; } 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; -x_42 = lean_ctor_get(x_38, 1); -lean_inc(x_42); -lean_dec(x_38); -x_43 = lean_ctor_get(x_39, 0); -lean_inc(x_43); -lean_dec(x_39); -x_44 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_44, 0, x_9); -lean_closure_set(x_44, 1, x_43); -x_45 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__2), 8, 3); -lean_closure_set(x_45, 0, x_2); -lean_closure_set(x_45, 1, x_1); -lean_closure_set(x_45, 2, x_3); -x_46 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_46, 0, x_44); -lean_closure_set(x_46, 1, x_45); -x_47 = l_Lean_Compiler_withNewScopeImp___rarg(x_46, x_5, x_6, x_7, x_42); -return x_47; -} -} -else +uint8_t x_1081; +x_1081 = lean_nat_dec_le(x_1078, x_1078); +if (x_1081 == 0) { -uint8_t x_48; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_1082; +lean_dec(x_1078); +lean_dec(x_1058); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_38); -if (x_48 == 0) +x_1082 = lean_box(0); +x_1003 = x_1082; +x_1004 = x_1060; +goto block_1021; +} +else { -return x_38; +size_t x_1083; size_t x_1084; lean_object* x_1085; lean_object* x_1086; +x_1083 = 0; +x_1084 = lean_usize_of_nat(x_1078); +lean_dec(x_1078); +x_1085 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1086 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__36(x_1, x_2, x_3, x_1058, x_1083, x_1084, x_1085, x_5, x_6, x_7, x_1060); +lean_dec(x_1058); +if (lean_obj_tag(x_1086) == 0) +{ +lean_object* x_1087; lean_object* x_1088; +x_1087 = lean_ctor_get(x_1086, 0); +lean_inc(x_1087); +x_1088 = lean_ctor_get(x_1086, 1); +lean_inc(x_1088); +lean_dec(x_1086); +x_1003 = x_1087; +x_1004 = x_1088; +goto block_1021; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_38, 0); -x_50 = lean_ctor_get(x_38, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_38); -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; +lean_object* x_1089; lean_object* x_1090; +x_1089 = lean_ctor_get(x_1086, 0); +lean_inc(x_1089); +x_1090 = lean_ctor_get(x_1086, 1); +lean_inc(x_1090); +lean_dec(x_1086); +x_1022 = x_1089; +x_1023 = x_1090; +goto block_1040; } } } -case 3: +} +else { -lean_object* x_52; -lean_dec(x_14); +size_t x_1091; size_t x_1092; lean_object* x_1093; lean_object* x_1094; +x_1091 = 0; +x_1092 = lean_usize_of_nat(x_1061); +lean_dec(x_1061); +x_1093 = lean_box(0); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_9); -x_52 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_52) == 0) +lean_inc(x_5); +lean_inc(x_2); +x_1094 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__37(x_2, x_1057, x_1091, x_1092, x_1093, x_5, x_6, x_7, x_1060); +lean_dec(x_1057); +if (lean_obj_tag(x_1094) == 0) { -lean_object* x_53; -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -if (lean_obj_tag(x_53) == 0) +lean_object* x_1095; lean_object* x_1096; uint8_t x_1097; +x_1095 = lean_ctor_get(x_1094, 1); +lean_inc(x_1095); +lean_dec(x_1094); +x_1096 = lean_array_get_size(x_1058); +x_1097 = lean_nat_dec_lt(x_1062, x_1096); +if (x_1097 == 0) { -lean_object* x_54; lean_object* x_55; +lean_dec(x_1096); +lean_dec(x_1058); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec(x_52); -x_55 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_54); -return x_55; +x_1003 = x_1093; +x_1004 = x_1095; +goto block_1021; } 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; -x_56 = lean_ctor_get(x_52, 1); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_ctor_get(x_53, 0); -lean_inc(x_57); -lean_dec(x_53); -x_58 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_58, 0, x_9); -lean_closure_set(x_58, 1, x_57); -x_59 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__3), 8, 3); -lean_closure_set(x_59, 0, x_2); -lean_closure_set(x_59, 1, x_1); -lean_closure_set(x_59, 2, x_3); -x_60 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_60, 0, x_58); -lean_closure_set(x_60, 1, x_59); -x_61 = l_Lean_Compiler_withNewScopeImp___rarg(x_60, x_5, x_6, x_7, x_56); -return x_61; -} -} -else +uint8_t x_1098; +x_1098 = lean_nat_dec_le(x_1096, x_1096); +if (x_1098 == 0) { -uint8_t x_62; -lean_dec(x_9); -lean_dec(x_7); +lean_dec(x_1096); +lean_dec(x_1058); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_62 = !lean_is_exclusive(x_52); -if (x_62 == 0) +x_1003 = x_1093; +x_1004 = x_1095; +goto block_1021; +} +else +{ +size_t x_1099; lean_object* x_1100; +x_1099 = lean_usize_of_nat(x_1096); +lean_dec(x_1096); +lean_inc(x_7); +lean_inc(x_5); +x_1100 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__38(x_1, x_2, x_3, x_1058, x_1091, x_1099, x_1093, x_5, x_6, x_7, x_1095); +lean_dec(x_1058); +if (lean_obj_tag(x_1100) == 0) { -return x_52; +lean_object* x_1101; lean_object* x_1102; +x_1101 = lean_ctor_get(x_1100, 0); +lean_inc(x_1101); +x_1102 = lean_ctor_get(x_1100, 1); +lean_inc(x_1102); +lean_dec(x_1100); +x_1003 = x_1101; +x_1004 = x_1102; +goto block_1021; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_52, 0); -x_64 = lean_ctor_get(x_52, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_52); -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; +lean_object* x_1103; lean_object* x_1104; +x_1103 = lean_ctor_get(x_1100, 0); +lean_inc(x_1103); +x_1104 = lean_ctor_get(x_1100, 1); +lean_inc(x_1104); +lean_dec(x_1100); +x_1022 = x_1103; +x_1023 = x_1104; +goto block_1040; } } } -case 4: -{ -lean_object* x_66; -lean_dec(x_14); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_66 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_66) == 0) +} +else { -lean_object* x_67; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -if (lean_obj_tag(x_67) == 0) +lean_object* x_1105; lean_object* x_1106; +lean_dec(x_1058); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1105 = lean_ctor_get(x_1094, 0); +lean_inc(x_1105); +x_1106 = lean_ctor_get(x_1094, 1); +lean_inc(x_1106); +lean_dec(x_1094); +x_1022 = x_1105; +x_1023 = x_1106; +goto block_1040; +} +} +} +} +else +{ +lean_object* x_1107; lean_object* x_1108; +lean_dec(x_1058); +lean_dec(x_1057); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1107 = lean_ctor_get(x_1059, 0); +lean_inc(x_1107); +x_1108 = lean_ctor_get(x_1059, 1); +lean_inc(x_1108); +lean_dec(x_1059); +x_1022 = x_1107; +x_1023 = x_1108; +goto block_1040; +} +block_1021: +{ +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; uint8_t x_1017; +x_1005 = lean_st_ref_get(x_7, x_1004); +x_1006 = lean_ctor_get(x_1005, 1); +lean_inc(x_1006); +lean_dec(x_1005); +x_1007 = lean_st_ref_get(x_5, x_1006); +x_1008 = lean_ctor_get(x_1007, 0); +lean_inc(x_1008); +x_1009 = lean_ctor_get(x_1007, 1); +lean_inc(x_1009); +lean_dec(x_1007); +x_1010 = lean_ctor_get(x_1001, 0); +lean_inc(x_1010); +x_1011 = lean_ctor_get(x_1001, 1); +lean_inc(x_1011); +lean_dec(x_1001); +x_1012 = lean_ctor_get(x_1008, 2); +lean_inc(x_1012); +lean_dec(x_1008); +x_1013 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1013, 0, x_1010); +lean_ctor_set(x_1013, 1, x_1011); +lean_ctor_set(x_1013, 2, x_1012); +x_1014 = lean_st_ref_get(x_7, x_1009); +lean_dec(x_7); +x_1015 = lean_ctor_get(x_1014, 1); +lean_inc(x_1015); +lean_dec(x_1014); +x_1016 = lean_st_ref_set(x_5, x_1013, x_1015); +lean_dec(x_5); +x_1017 = !lean_is_exclusive(x_1016); +if (x_1017 == 0) { -lean_object* x_68; lean_object* x_69; -lean_dec(x_3); -lean_dec(x_1); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_68); -return x_69; +lean_object* x_1018; +x_1018 = lean_ctor_get(x_1016, 0); +lean_dec(x_1018); +lean_ctor_set(x_1016, 0, x_1003); +return x_1016; } 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_66, 1); -lean_inc(x_70); -lean_dec(x_66); -x_71 = lean_ctor_get(x_67, 0); -lean_inc(x_71); -lean_dec(x_67); -x_72 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_72, 0, x_9); -lean_closure_set(x_72, 1, x_71); -x_73 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__4), 8, 3); -lean_closure_set(x_73, 0, x_2); -lean_closure_set(x_73, 1, x_1); -lean_closure_set(x_73, 2, x_3); -x_74 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_74, 0, x_72); -lean_closure_set(x_74, 1, x_73); -x_75 = l_Lean_Compiler_withNewScopeImp___rarg(x_74, x_5, x_6, x_7, x_70); -return x_75; +lean_object* x_1019; lean_object* x_1020; +x_1019 = lean_ctor_get(x_1016, 1); +lean_inc(x_1019); +lean_dec(x_1016); +x_1020 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1020, 0, x_1003); +lean_ctor_set(x_1020, 1, x_1019); +return x_1020; } } -else +block_1040: { -uint8_t x_76; +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; uint8_t x_1036; +x_1024 = lean_st_ref_get(x_7, x_1023); +x_1025 = lean_ctor_get(x_1024, 1); +lean_inc(x_1025); +lean_dec(x_1024); +x_1026 = lean_st_ref_get(x_5, x_1025); +x_1027 = lean_ctor_get(x_1026, 0); +lean_inc(x_1027); +x_1028 = lean_ctor_get(x_1026, 1); +lean_inc(x_1028); +lean_dec(x_1026); +x_1029 = lean_ctor_get(x_1001, 0); +lean_inc(x_1029); +x_1030 = lean_ctor_get(x_1001, 1); +lean_inc(x_1030); +lean_dec(x_1001); +x_1031 = lean_ctor_get(x_1027, 2); +lean_inc(x_1031); +lean_dec(x_1027); +x_1032 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1032, 0, x_1029); +lean_ctor_set(x_1032, 1, x_1030); +lean_ctor_set(x_1032, 2, x_1031); +x_1033 = lean_st_ref_get(x_7, x_1028); +lean_dec(x_7); +x_1034 = lean_ctor_get(x_1033, 1); +lean_inc(x_1034); +lean_dec(x_1033); +x_1035 = lean_st_ref_set(x_5, x_1032, x_1034); +lean_dec(x_5); +x_1036 = !lean_is_exclusive(x_1035); +if (x_1036 == 0) +{ +lean_object* x_1037; +x_1037 = lean_ctor_get(x_1035, 0); +lean_dec(x_1037); +lean_ctor_set_tag(x_1035, 1); +lean_ctor_set(x_1035, 0, x_1022); +return x_1035; +} +else +{ +lean_object* x_1038; lean_object* x_1039; +x_1038 = lean_ctor_get(x_1035, 1); +lean_inc(x_1038); +lean_dec(x_1035); +x_1039 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1039, 0, x_1022); +lean_ctor_set(x_1039, 1, x_1038); +return x_1039; +} +} +} +} +else +{ +uint8_t x_1109; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -30907,307 +30115,516 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_76 = !lean_is_exclusive(x_66); -if (x_76 == 0) +x_1109 = !lean_is_exclusive(x_992); +if (x_1109 == 0) { -return x_66; +return x_992; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_66, 0); -x_78 = lean_ctor_get(x_66, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_66); -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; +lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; +x_1110 = lean_ctor_get(x_992, 0); +x_1111 = lean_ctor_get(x_992, 1); +lean_inc(x_1111); +lean_inc(x_1110); +lean_dec(x_992); +x_1112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1112, 0, x_1110); +lean_ctor_set(x_1112, 1, x_1111); +return x_1112; } } } -case 5: +case 10: { -lean_object* x_80; +lean_object* x_1113; lean_dec(x_14); lean_inc(x_7); lean_inc(x_6); lean_inc(x_9); -x_80 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_80) == 0) +x_1113 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_1113) == 0) +{ +lean_object* x_1114; +x_1114 = lean_ctor_get(x_1113, 0); +lean_inc(x_1114); +if (lean_obj_tag(x_1114) == 0) +{ +lean_object* x_1115; lean_object* x_1116; +lean_dec(x_3); +lean_dec(x_1); +x_1115 = lean_ctor_get(x_1113, 1); +lean_inc(x_1115); +lean_dec(x_1113); +x_1116 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_1115); +return x_1116; +} +else +{ +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_1143; lean_object* x_1144; 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; +x_1117 = lean_ctor_get(x_1113, 1); +lean_inc(x_1117); +lean_dec(x_1113); +x_1118 = lean_ctor_get(x_1114, 0); +lean_inc(x_1118); +lean_dec(x_1114); +x_1119 = lean_st_ref_get(x_7, x_1117); +x_1120 = lean_ctor_get(x_1119, 1); +lean_inc(x_1120); +lean_dec(x_1119); +x_1121 = lean_st_ref_get(x_5, x_1120); +x_1122 = lean_ctor_get(x_1121, 0); +lean_inc(x_1122); +x_1123 = lean_ctor_get(x_1121, 1); +lean_inc(x_1123); +lean_dec(x_1121); +x_1162 = lean_st_ref_get(x_7, x_1123); +x_1163 = lean_ctor_get(x_1162, 1); +lean_inc(x_1163); +lean_dec(x_1162); +x_1164 = lean_st_ref_take(x_5, x_1163); +x_1165 = lean_ctor_get(x_1164, 0); +lean_inc(x_1165); +x_1166 = lean_ctor_get(x_1164, 1); +lean_inc(x_1166); +lean_dec(x_1164); +x_1167 = lean_ctor_get(x_1165, 0); +lean_inc(x_1167); +x_1168 = lean_ctor_get(x_1165, 2); +lean_inc(x_1168); +lean_dec(x_1165); +x_1169 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1170, 0, x_1167); +lean_ctor_set(x_1170, 1, x_1169); +lean_ctor_set(x_1170, 2, x_1168); +x_1171 = lean_st_ref_set(x_5, x_1170, x_1166); +x_1172 = lean_ctor_get(x_1171, 1); +lean_inc(x_1172); +lean_dec(x_1171); +x_1173 = l_Lean_Compiler_visitMatch(x_9, x_1118, x_5, x_6, x_7, x_1172); +x_1174 = lean_ctor_get(x_1173, 0); +lean_inc(x_1174); +x_1175 = lean_ctor_get(x_1174, 1); +lean_inc(x_1175); +x_1176 = lean_ctor_get(x_1173, 1); +lean_inc(x_1176); +lean_dec(x_1173); +x_1177 = lean_ctor_get(x_1174, 0); +lean_inc(x_1177); +lean_dec(x_1174); +x_1178 = lean_ctor_get(x_1175, 0); +lean_inc(x_1178); +x_1179 = lean_ctor_get(x_1175, 1); +lean_inc(x_1179); +lean_dec(x_1175); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_1180 = lean_apply_5(x_2, x_1177, x_5, x_6, x_7, x_1176); +if (lean_obj_tag(x_1180) == 0) { -lean_object* x_81; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -if (lean_obj_tag(x_81) == 0) +lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; uint8_t x_1184; +x_1181 = lean_ctor_get(x_1180, 1); +lean_inc(x_1181); +lean_dec(x_1180); +x_1182 = lean_array_get_size(x_1178); +x_1183 = lean_unsigned_to_nat(0u); +x_1184 = lean_nat_dec_lt(x_1183, x_1182); +if (x_1184 == 0) +{ +lean_object* x_1185; uint8_t x_1186; +lean_dec(x_1182); +lean_dec(x_1178); +x_1185 = lean_array_get_size(x_1179); +x_1186 = lean_nat_dec_lt(x_1183, x_1185); +if (x_1186 == 0) { -lean_object* x_82; lean_object* x_83; +lean_object* x_1187; +lean_dec(x_1185); +lean_dec(x_1179); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -lean_dec(x_80); -x_83 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_82); -return x_83; +x_1187 = lean_box(0); +x_1124 = x_1187; +x_1125 = x_1181; +goto block_1142; } 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_80, 1); -lean_inc(x_84); -lean_dec(x_80); -x_85 = lean_ctor_get(x_81, 0); -lean_inc(x_85); -lean_dec(x_81); -x_86 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_86, 0, x_9); -lean_closure_set(x_86, 1, x_85); -x_87 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__5), 8, 3); -lean_closure_set(x_87, 0, x_2); -lean_closure_set(x_87, 1, x_1); -lean_closure_set(x_87, 2, x_3); -x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_88, 0, x_86); -lean_closure_set(x_88, 1, x_87); -x_89 = l_Lean_Compiler_withNewScopeImp___rarg(x_88, x_5, x_6, x_7, x_84); -return x_89; -} -} -else +uint8_t x_1188; +x_1188 = lean_nat_dec_le(x_1185, x_1185); +if (x_1188 == 0) { -uint8_t x_90; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_1189; +lean_dec(x_1185); +lean_dec(x_1179); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_90 = !lean_is_exclusive(x_80); -if (x_90 == 0) +x_1189 = lean_box(0); +x_1124 = x_1189; +x_1125 = x_1181; +goto block_1142; +} +else { -return x_80; +size_t x_1190; size_t x_1191; lean_object* x_1192; lean_object* x_1193; +x_1190 = 0; +x_1191 = lean_usize_of_nat(x_1185); +lean_dec(x_1185); +x_1192 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1193 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__39(x_1, x_2, x_3, x_1179, x_1190, x_1191, x_1192, x_5, x_6, x_7, x_1181); +lean_dec(x_1179); +if (lean_obj_tag(x_1193) == 0) +{ +lean_object* x_1194; lean_object* x_1195; +x_1194 = lean_ctor_get(x_1193, 0); +lean_inc(x_1194); +x_1195 = lean_ctor_get(x_1193, 1); +lean_inc(x_1195); +lean_dec(x_1193); +x_1124 = x_1194; +x_1125 = x_1195; +goto block_1142; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_80, 0); -x_92 = lean_ctor_get(x_80, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_80); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_object* x_1196; lean_object* x_1197; +x_1196 = lean_ctor_get(x_1193, 0); +lean_inc(x_1196); +x_1197 = lean_ctor_get(x_1193, 1); +lean_inc(x_1197); +lean_dec(x_1193); +x_1143 = x_1196; +x_1144 = x_1197; +goto block_1161; } } } -case 6: +} +else { -lean_object* x_94; -lean_dec(x_14); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_94 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_94) == 0) +uint8_t x_1198; +x_1198 = lean_nat_dec_le(x_1182, x_1182); +if (x_1198 == 0) { -lean_object* x_95; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -if (lean_obj_tag(x_95) == 0) +lean_object* x_1199; uint8_t x_1200; +lean_dec(x_1182); +lean_dec(x_1178); +x_1199 = lean_array_get_size(x_1179); +x_1200 = lean_nat_dec_lt(x_1183, x_1199); +if (x_1200 == 0) { -lean_object* x_96; lean_object* x_97; +lean_object* x_1201; +lean_dec(x_1199); +lean_dec(x_1179); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_96); -return x_97; +x_1201 = lean_box(0); +x_1124 = x_1201; +x_1125 = x_1181; +goto block_1142; } else { -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_94, 1); -lean_inc(x_98); -lean_dec(x_94); -x_99 = lean_ctor_get(x_95, 0); -lean_inc(x_99); -lean_dec(x_95); -x_100 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_100, 0, x_9); -lean_closure_set(x_100, 1, x_99); -x_101 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__6), 8, 3); -lean_closure_set(x_101, 0, x_2); -lean_closure_set(x_101, 1, x_1); -lean_closure_set(x_101, 2, x_3); -x_102 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_102, 0, x_100); -lean_closure_set(x_102, 1, x_101); -x_103 = l_Lean_Compiler_withNewScopeImp___rarg(x_102, x_5, x_6, x_7, x_98); -return x_103; -} -} -else +uint8_t x_1202; +x_1202 = lean_nat_dec_le(x_1199, x_1199); +if (x_1202 == 0) { -uint8_t x_104; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_1203; +lean_dec(x_1199); +lean_dec(x_1179); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_104 = !lean_is_exclusive(x_94); -if (x_104 == 0) +x_1203 = lean_box(0); +x_1124 = x_1203; +x_1125 = x_1181; +goto block_1142; +} +else +{ +size_t x_1204; size_t x_1205; lean_object* x_1206; lean_object* x_1207; +x_1204 = 0; +x_1205 = lean_usize_of_nat(x_1199); +lean_dec(x_1199); +x_1206 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1207 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__40(x_1, x_2, x_3, x_1179, x_1204, x_1205, x_1206, x_5, x_6, x_7, x_1181); +lean_dec(x_1179); +if (lean_obj_tag(x_1207) == 0) { -return x_94; +lean_object* x_1208; lean_object* x_1209; +x_1208 = lean_ctor_get(x_1207, 0); +lean_inc(x_1208); +x_1209 = lean_ctor_get(x_1207, 1); +lean_inc(x_1209); +lean_dec(x_1207); +x_1124 = x_1208; +x_1125 = x_1209; +goto block_1142; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_94, 0); -x_106 = lean_ctor_get(x_94, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_94); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -return x_107; +lean_object* x_1210; lean_object* x_1211; +x_1210 = lean_ctor_get(x_1207, 0); +lean_inc(x_1210); +x_1211 = lean_ctor_get(x_1207, 1); +lean_inc(x_1211); +lean_dec(x_1207); +x_1143 = x_1210; +x_1144 = x_1211; +goto block_1161; } } } -case 7: +} +else { -lean_object* x_108; -lean_dec(x_14); +size_t x_1212; size_t x_1213; lean_object* x_1214; lean_object* x_1215; +x_1212 = 0; +x_1213 = lean_usize_of_nat(x_1182); +lean_dec(x_1182); +x_1214 = lean_box(0); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_9); -x_108 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_108) == 0) +lean_inc(x_5); +lean_inc(x_2); +x_1215 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__41(x_2, x_1178, x_1212, x_1213, x_1214, x_5, x_6, x_7, x_1181); +lean_dec(x_1178); +if (lean_obj_tag(x_1215) == 0) { -lean_object* x_109; -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_obj_tag(x_109) == 0) +lean_object* x_1216; lean_object* x_1217; uint8_t x_1218; +x_1216 = lean_ctor_get(x_1215, 1); +lean_inc(x_1216); +lean_dec(x_1215); +x_1217 = lean_array_get_size(x_1179); +x_1218 = lean_nat_dec_lt(x_1183, x_1217); +if (x_1218 == 0) { -lean_object* x_110; lean_object* x_111; +lean_dec(x_1217); +lean_dec(x_1179); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_110); -return x_111; +x_1124 = x_1214; +x_1125 = x_1216; +goto block_1142; } 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; -x_112 = lean_ctor_get(x_108, 1); -lean_inc(x_112); -lean_dec(x_108); -x_113 = lean_ctor_get(x_109, 0); -lean_inc(x_113); -lean_dec(x_109); -x_114 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_114, 0, x_9); -lean_closure_set(x_114, 1, x_113); -x_115 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__7), 8, 3); -lean_closure_set(x_115, 0, x_2); -lean_closure_set(x_115, 1, x_1); -lean_closure_set(x_115, 2, x_3); -x_116 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_116, 0, x_114); -lean_closure_set(x_116, 1, x_115); -x_117 = l_Lean_Compiler_withNewScopeImp___rarg(x_116, x_5, x_6, x_7, x_112); -return x_117; -} -} -else +uint8_t x_1219; +x_1219 = lean_nat_dec_le(x_1217, x_1217); +if (x_1219 == 0) { -uint8_t x_118; -lean_dec(x_9); -lean_dec(x_7); +lean_dec(x_1217); +lean_dec(x_1179); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_118 = !lean_is_exclusive(x_108); -if (x_118 == 0) +x_1124 = x_1214; +x_1125 = x_1216; +goto block_1142; +} +else +{ +size_t x_1220; lean_object* x_1221; +x_1220 = lean_usize_of_nat(x_1217); +lean_dec(x_1217); +lean_inc(x_7); +lean_inc(x_5); +x_1221 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__42(x_1, x_2, x_3, x_1179, x_1212, x_1220, x_1214, x_5, x_6, x_7, x_1216); +lean_dec(x_1179); +if (lean_obj_tag(x_1221) == 0) { -return x_108; +lean_object* x_1222; lean_object* x_1223; +x_1222 = lean_ctor_get(x_1221, 0); +lean_inc(x_1222); +x_1223 = lean_ctor_get(x_1221, 1); +lean_inc(x_1223); +lean_dec(x_1221); +x_1124 = x_1222; +x_1125 = x_1223; +goto block_1142; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_108, 0); -x_120 = lean_ctor_get(x_108, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_108); -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; +lean_object* x_1224; lean_object* x_1225; +x_1224 = lean_ctor_get(x_1221, 0); +lean_inc(x_1224); +x_1225 = lean_ctor_get(x_1221, 1); +lean_inc(x_1225); +lean_dec(x_1221); +x_1143 = x_1224; +x_1144 = x_1225; +goto block_1161; } } } -case 8: -{ -lean_object* x_122; -lean_dec(x_14); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_122 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_122) == 0) -{ -lean_object* x_123; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -if (lean_obj_tag(x_123) == 0) +} +else { -lean_object* x_124; lean_object* x_125; +lean_object* x_1226; lean_object* x_1227; +lean_dec(x_1179); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_124); -return x_125; +x_1226 = lean_ctor_get(x_1215, 0); +lean_inc(x_1226); +x_1227 = lean_ctor_get(x_1215, 1); +lean_inc(x_1227); +lean_dec(x_1215); +x_1143 = x_1226; +x_1144 = x_1227; +goto block_1161; +} +} +} +} +else +{ +lean_object* x_1228; lean_object* x_1229; +lean_dec(x_1179); +lean_dec(x_1178); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1228 = lean_ctor_get(x_1180, 0); +lean_inc(x_1228); +x_1229 = lean_ctor_get(x_1180, 1); +lean_inc(x_1229); +lean_dec(x_1180); +x_1143 = x_1228; +x_1144 = x_1229; +goto block_1161; +} +block_1142: +{ +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; uint8_t x_1138; +x_1126 = lean_st_ref_get(x_7, x_1125); +x_1127 = lean_ctor_get(x_1126, 1); +lean_inc(x_1127); +lean_dec(x_1126); +x_1128 = lean_st_ref_get(x_5, x_1127); +x_1129 = lean_ctor_get(x_1128, 0); +lean_inc(x_1129); +x_1130 = lean_ctor_get(x_1128, 1); +lean_inc(x_1130); +lean_dec(x_1128); +x_1131 = lean_ctor_get(x_1122, 0); +lean_inc(x_1131); +x_1132 = lean_ctor_get(x_1122, 1); +lean_inc(x_1132); +lean_dec(x_1122); +x_1133 = lean_ctor_get(x_1129, 2); +lean_inc(x_1133); +lean_dec(x_1129); +x_1134 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1134, 0, x_1131); +lean_ctor_set(x_1134, 1, x_1132); +lean_ctor_set(x_1134, 2, x_1133); +x_1135 = lean_st_ref_get(x_7, x_1130); +lean_dec(x_7); +x_1136 = lean_ctor_get(x_1135, 1); +lean_inc(x_1136); +lean_dec(x_1135); +x_1137 = lean_st_ref_set(x_5, x_1134, x_1136); +lean_dec(x_5); +x_1138 = !lean_is_exclusive(x_1137); +if (x_1138 == 0) +{ +lean_object* x_1139; +x_1139 = lean_ctor_get(x_1137, 0); +lean_dec(x_1139); +lean_ctor_set(x_1137, 0, x_1124); +return x_1137; } 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; -x_126 = lean_ctor_get(x_122, 1); -lean_inc(x_126); -lean_dec(x_122); -x_127 = lean_ctor_get(x_123, 0); -lean_inc(x_127); -lean_dec(x_123); -x_128 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_128, 0, x_9); -lean_closure_set(x_128, 1, x_127); -x_129 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__8), 8, 3); -lean_closure_set(x_129, 0, x_2); -lean_closure_set(x_129, 1, x_1); -lean_closure_set(x_129, 2, x_3); -x_130 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_130, 0, x_128); -lean_closure_set(x_130, 1, x_129); -x_131 = l_Lean_Compiler_withNewScopeImp___rarg(x_130, x_5, x_6, x_7, x_126); -return x_131; +lean_object* x_1140; lean_object* x_1141; +x_1140 = lean_ctor_get(x_1137, 1); +lean_inc(x_1140); +lean_dec(x_1137); +x_1141 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1141, 0, x_1124); +lean_ctor_set(x_1141, 1, x_1140); +return x_1141; } } -else +block_1161: { -uint8_t x_132; +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; uint8_t x_1157; +x_1145 = lean_st_ref_get(x_7, x_1144); +x_1146 = lean_ctor_get(x_1145, 1); +lean_inc(x_1146); +lean_dec(x_1145); +x_1147 = lean_st_ref_get(x_5, x_1146); +x_1148 = lean_ctor_get(x_1147, 0); +lean_inc(x_1148); +x_1149 = lean_ctor_get(x_1147, 1); +lean_inc(x_1149); +lean_dec(x_1147); +x_1150 = lean_ctor_get(x_1122, 0); +lean_inc(x_1150); +x_1151 = lean_ctor_get(x_1122, 1); +lean_inc(x_1151); +lean_dec(x_1122); +x_1152 = lean_ctor_get(x_1148, 2); +lean_inc(x_1152); +lean_dec(x_1148); +x_1153 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1153, 0, x_1150); +lean_ctor_set(x_1153, 1, x_1151); +lean_ctor_set(x_1153, 2, x_1152); +x_1154 = lean_st_ref_get(x_7, x_1149); +lean_dec(x_7); +x_1155 = lean_ctor_get(x_1154, 1); +lean_inc(x_1155); +lean_dec(x_1154); +x_1156 = lean_st_ref_set(x_5, x_1153, x_1155); +lean_dec(x_5); +x_1157 = !lean_is_exclusive(x_1156); +if (x_1157 == 0) +{ +lean_object* x_1158; +x_1158 = lean_ctor_get(x_1156, 0); +lean_dec(x_1158); +lean_ctor_set_tag(x_1156, 1); +lean_ctor_set(x_1156, 0, x_1143); +return x_1156; +} +else +{ +lean_object* x_1159; lean_object* x_1160; +x_1159 = lean_ctor_get(x_1156, 1); +lean_inc(x_1159); +lean_dec(x_1156); +x_1160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1160, 0, x_1143); +lean_ctor_set(x_1160, 1, x_1159); +return x_1160; +} +} +} +} +else +{ +uint8_t x_1230; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -31215,230 +30632,516 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_132 = !lean_is_exclusive(x_122); -if (x_132 == 0) +x_1230 = !lean_is_exclusive(x_1113); +if (x_1230 == 0) { -return x_122; +return x_1113; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_122, 0); -x_134 = lean_ctor_get(x_122, 1); -lean_inc(x_134); -lean_inc(x_133); -lean_dec(x_122); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -return x_135; +lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; +x_1231 = lean_ctor_get(x_1113, 0); +x_1232 = lean_ctor_get(x_1113, 1); +lean_inc(x_1232); +lean_inc(x_1231); +lean_dec(x_1113); +x_1233 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1233, 0, x_1231); +lean_ctor_set(x_1233, 1, x_1232); +return x_1233; } } } -case 9: +default: { -lean_object* x_136; +lean_object* x_1234; lean_dec(x_14); lean_inc(x_7); lean_inc(x_6); lean_inc(x_9); -x_136 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_136) == 0) +x_1234 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); +if (lean_obj_tag(x_1234) == 0) +{ +lean_object* x_1235; +x_1235 = lean_ctor_get(x_1234, 0); +lean_inc(x_1235); +if (lean_obj_tag(x_1235) == 0) +{ +lean_object* x_1236; lean_object* x_1237; +lean_dec(x_3); +lean_dec(x_1); +x_1236 = lean_ctor_get(x_1234, 1); +lean_inc(x_1236); +lean_dec(x_1234); +x_1237 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_1236); +return x_1237; +} +else +{ +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_1264; lean_object* x_1265; 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; +x_1238 = lean_ctor_get(x_1234, 1); +lean_inc(x_1238); +lean_dec(x_1234); +x_1239 = lean_ctor_get(x_1235, 0); +lean_inc(x_1239); +lean_dec(x_1235); +x_1240 = lean_st_ref_get(x_7, x_1238); +x_1241 = lean_ctor_get(x_1240, 1); +lean_inc(x_1241); +lean_dec(x_1240); +x_1242 = lean_st_ref_get(x_5, x_1241); +x_1243 = lean_ctor_get(x_1242, 0); +lean_inc(x_1243); +x_1244 = lean_ctor_get(x_1242, 1); +lean_inc(x_1244); +lean_dec(x_1242); +x_1283 = lean_st_ref_get(x_7, x_1244); +x_1284 = lean_ctor_get(x_1283, 1); +lean_inc(x_1284); +lean_dec(x_1283); +x_1285 = lean_st_ref_take(x_5, x_1284); +x_1286 = lean_ctor_get(x_1285, 0); +lean_inc(x_1286); +x_1287 = lean_ctor_get(x_1285, 1); +lean_inc(x_1287); +lean_dec(x_1285); +x_1288 = lean_ctor_get(x_1286, 0); +lean_inc(x_1288); +x_1289 = lean_ctor_get(x_1286, 2); +lean_inc(x_1289); +lean_dec(x_1286); +x_1290 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1291 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1291, 0, x_1288); +lean_ctor_set(x_1291, 1, x_1290); +lean_ctor_set(x_1291, 2, x_1289); +x_1292 = lean_st_ref_set(x_5, x_1291, x_1287); +x_1293 = lean_ctor_get(x_1292, 1); +lean_inc(x_1293); +lean_dec(x_1292); +x_1294 = l_Lean_Compiler_visitMatch(x_9, x_1239, x_5, x_6, x_7, x_1293); +x_1295 = lean_ctor_get(x_1294, 0); +lean_inc(x_1295); +x_1296 = lean_ctor_get(x_1295, 1); +lean_inc(x_1296); +x_1297 = lean_ctor_get(x_1294, 1); +lean_inc(x_1297); +lean_dec(x_1294); +x_1298 = lean_ctor_get(x_1295, 0); +lean_inc(x_1298); +lean_dec(x_1295); +x_1299 = lean_ctor_get(x_1296, 0); +lean_inc(x_1299); +x_1300 = lean_ctor_get(x_1296, 1); +lean_inc(x_1300); +lean_dec(x_1296); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_1301 = lean_apply_5(x_2, x_1298, x_5, x_6, x_7, x_1297); +if (lean_obj_tag(x_1301) == 0) { -lean_object* x_137; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -if (lean_obj_tag(x_137) == 0) +lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; uint8_t x_1305; +x_1302 = lean_ctor_get(x_1301, 1); +lean_inc(x_1302); +lean_dec(x_1301); +x_1303 = lean_array_get_size(x_1299); +x_1304 = lean_unsigned_to_nat(0u); +x_1305 = lean_nat_dec_lt(x_1304, x_1303); +if (x_1305 == 0) { -lean_object* x_138; lean_object* x_139; +lean_object* x_1306; uint8_t x_1307; +lean_dec(x_1303); +lean_dec(x_1299); +x_1306 = lean_array_get_size(x_1300); +x_1307 = lean_nat_dec_lt(x_1304, x_1306); +if (x_1307 == 0) +{ +lean_object* x_1308; +lean_dec(x_1306); +lean_dec(x_1300); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -x_139 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_138); -return x_139; +x_1308 = lean_box(0); +x_1245 = x_1308; +x_1246 = x_1302; +goto block_1263; } 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; -x_140 = lean_ctor_get(x_136, 1); -lean_inc(x_140); -lean_dec(x_136); -x_141 = lean_ctor_get(x_137, 0); -lean_inc(x_141); -lean_dec(x_137); -x_142 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_142, 0, x_9); -lean_closure_set(x_142, 1, x_141); -x_143 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__9), 8, 3); -lean_closure_set(x_143, 0, x_2); -lean_closure_set(x_143, 1, x_1); -lean_closure_set(x_143, 2, x_3); -x_144 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_144, 0, x_142); -lean_closure_set(x_144, 1, x_143); -x_145 = l_Lean_Compiler_withNewScopeImp___rarg(x_144, x_5, x_6, x_7, x_140); -return x_145; -} -} -else +uint8_t x_1309; +x_1309 = lean_nat_dec_le(x_1306, x_1306); +if (x_1309 == 0) { -uint8_t x_146; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_1310; +lean_dec(x_1306); +lean_dec(x_1300); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_146 = !lean_is_exclusive(x_136); -if (x_146 == 0) +x_1310 = lean_box(0); +x_1245 = x_1310; +x_1246 = x_1302; +goto block_1263; +} +else { -return x_136; +size_t x_1311; size_t x_1312; lean_object* x_1313; lean_object* x_1314; +x_1311 = 0; +x_1312 = lean_usize_of_nat(x_1306); +lean_dec(x_1306); +x_1313 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1314 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__43(x_1, x_2, x_3, x_1300, x_1311, x_1312, x_1313, x_5, x_6, x_7, x_1302); +lean_dec(x_1300); +if (lean_obj_tag(x_1314) == 0) +{ +lean_object* x_1315; lean_object* x_1316; +x_1315 = lean_ctor_get(x_1314, 0); +lean_inc(x_1315); +x_1316 = lean_ctor_get(x_1314, 1); +lean_inc(x_1316); +lean_dec(x_1314); +x_1245 = x_1315; +x_1246 = x_1316; +goto block_1263; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_136, 0); -x_148 = lean_ctor_get(x_136, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_136); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_object* x_1317; lean_object* x_1318; +x_1317 = lean_ctor_get(x_1314, 0); +lean_inc(x_1317); +x_1318 = lean_ctor_get(x_1314, 1); +lean_inc(x_1318); +lean_dec(x_1314); +x_1264 = x_1317; +x_1265 = x_1318; +goto block_1282; } } } -case 10: +} +else { -lean_object* x_150; -lean_dec(x_14); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_9); -x_150 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_150) == 0) +uint8_t x_1319; +x_1319 = lean_nat_dec_le(x_1303, x_1303); +if (x_1319 == 0) { -lean_object* x_151; -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -if (lean_obj_tag(x_151) == 0) +lean_object* x_1320; uint8_t x_1321; +lean_dec(x_1303); +lean_dec(x_1299); +x_1320 = lean_array_get_size(x_1300); +x_1321 = lean_nat_dec_lt(x_1304, x_1320); +if (x_1321 == 0) { -lean_object* x_152; lean_object* x_153; +lean_object* x_1322; +lean_dec(x_1320); +lean_dec(x_1300); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_153 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_152); -return x_153; +x_1322 = lean_box(0); +x_1245 = x_1322; +x_1246 = x_1302; +goto block_1263; } else { -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_154 = lean_ctor_get(x_150, 1); -lean_inc(x_154); -lean_dec(x_150); -x_155 = lean_ctor_get(x_151, 0); -lean_inc(x_155); -lean_dec(x_151); -x_156 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_156, 0, x_9); -lean_closure_set(x_156, 1, x_155); -x_157 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__10), 8, 3); -lean_closure_set(x_157, 0, x_2); -lean_closure_set(x_157, 1, x_1); -lean_closure_set(x_157, 2, x_3); -x_158 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_158, 0, x_156); -lean_closure_set(x_158, 1, x_157); -x_159 = l_Lean_Compiler_withNewScopeImp___rarg(x_158, x_5, x_6, x_7, x_154); -return x_159; -} -} -else +uint8_t x_1323; +x_1323 = lean_nat_dec_le(x_1320, x_1320); +if (x_1323 == 0) { -uint8_t x_160; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_1324; +lean_dec(x_1320); +lean_dec(x_1300); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_160 = !lean_is_exclusive(x_150); -if (x_160 == 0) +x_1324 = lean_box(0); +x_1245 = x_1324; +x_1246 = x_1302; +goto block_1263; +} +else { -return x_150; +size_t x_1325; size_t x_1326; lean_object* x_1327; lean_object* x_1328; +x_1325 = 0; +x_1326 = lean_usize_of_nat(x_1320); +lean_dec(x_1320); +x_1327 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_1328 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__44(x_1, x_2, x_3, x_1300, x_1325, x_1326, x_1327, x_5, x_6, x_7, x_1302); +lean_dec(x_1300); +if (lean_obj_tag(x_1328) == 0) +{ +lean_object* x_1329; lean_object* x_1330; +x_1329 = lean_ctor_get(x_1328, 0); +lean_inc(x_1329); +x_1330 = lean_ctor_get(x_1328, 1); +lean_inc(x_1330); +lean_dec(x_1328); +x_1245 = x_1329; +x_1246 = x_1330; +goto block_1263; } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_150, 0); -x_162 = lean_ctor_get(x_150, 1); -lean_inc(x_162); -lean_inc(x_161); -lean_dec(x_150); -x_163 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_163, 0, x_161); -lean_ctor_set(x_163, 1, x_162); -return x_163; +lean_object* x_1331; lean_object* x_1332; +x_1331 = lean_ctor_get(x_1328, 0); +lean_inc(x_1331); +x_1332 = lean_ctor_get(x_1328, 1); +lean_inc(x_1332); +lean_dec(x_1328); +x_1264 = x_1331; +x_1265 = x_1332; +goto block_1282; } } } -default: +} +else { -lean_object* x_164; -lean_dec(x_14); +size_t x_1333; size_t x_1334; lean_object* x_1335; lean_object* x_1336; +x_1333 = 0; +x_1334 = lean_usize_of_nat(x_1303); +lean_dec(x_1303); +x_1335 = lean_box(0); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_9); -x_164 = l_Lean_Compiler_isCasesApp_x3f(x_9, x_6, x_7, x_8); -if (lean_obj_tag(x_164) == 0) +lean_inc(x_5); +lean_inc(x_2); +x_1336 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__45(x_2, x_1299, x_1333, x_1334, x_1335, x_5, x_6, x_7, x_1302); +lean_dec(x_1299); +if (lean_obj_tag(x_1336) == 0) +{ +lean_object* x_1337; lean_object* x_1338; uint8_t x_1339; +x_1337 = lean_ctor_get(x_1336, 1); +lean_inc(x_1337); +lean_dec(x_1336); +x_1338 = lean_array_get_size(x_1300); +x_1339 = lean_nat_dec_lt(x_1304, x_1338); +if (x_1339 == 0) +{ +lean_dec(x_1338); +lean_dec(x_1300); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1245 = x_1335; +x_1246 = x_1337; +goto block_1263; +} +else +{ +uint8_t x_1340; +x_1340 = lean_nat_dec_le(x_1338, x_1338); +if (x_1340 == 0) +{ +lean_dec(x_1338); +lean_dec(x_1300); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1245 = x_1335; +x_1246 = x_1337; +goto block_1263; +} +else +{ +size_t x_1341; lean_object* x_1342; +x_1341 = lean_usize_of_nat(x_1338); +lean_dec(x_1338); +lean_inc(x_7); +lean_inc(x_5); +x_1342 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__46(x_1, x_2, x_3, x_1300, x_1333, x_1341, x_1335, x_5, x_6, x_7, x_1337); +lean_dec(x_1300); +if (lean_obj_tag(x_1342) == 0) +{ +lean_object* x_1343; lean_object* x_1344; +x_1343 = lean_ctor_get(x_1342, 0); +lean_inc(x_1343); +x_1344 = lean_ctor_get(x_1342, 1); +lean_inc(x_1344); +lean_dec(x_1342); +x_1245 = x_1343; +x_1246 = x_1344; +goto block_1263; +} +else { -lean_object* x_165; -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -if (lean_obj_tag(x_165) == 0) +lean_object* x_1345; lean_object* x_1346; +x_1345 = lean_ctor_get(x_1342, 0); +lean_inc(x_1345); +x_1346 = lean_ctor_get(x_1342, 1); +lean_inc(x_1346); +lean_dec(x_1342); +x_1264 = x_1345; +x_1265 = x_1346; +goto block_1282; +} +} +} +} +else { -lean_object* x_166; lean_object* x_167; +lean_object* x_1347; lean_object* x_1348; +lean_dec(x_1300); +lean_dec(x_6); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_166); -return x_167; +x_1347 = lean_ctor_get(x_1336, 0); +lean_inc(x_1347); +x_1348 = lean_ctor_get(x_1336, 1); +lean_inc(x_1348); +lean_dec(x_1336); +x_1264 = x_1347; +x_1265 = x_1348; +goto block_1282; +} +} +} +} +else +{ +lean_object* x_1349; lean_object* x_1350; +lean_dec(x_1300); +lean_dec(x_1299); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1349 = lean_ctor_get(x_1301, 0); +lean_inc(x_1349); +x_1350 = lean_ctor_get(x_1301, 1); +lean_inc(x_1350); +lean_dec(x_1301); +x_1264 = x_1349; +x_1265 = x_1350; +goto block_1282; +} +block_1263: +{ +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; uint8_t x_1259; +x_1247 = lean_st_ref_get(x_7, x_1246); +x_1248 = lean_ctor_get(x_1247, 1); +lean_inc(x_1248); +lean_dec(x_1247); +x_1249 = lean_st_ref_get(x_5, x_1248); +x_1250 = lean_ctor_get(x_1249, 0); +lean_inc(x_1250); +x_1251 = lean_ctor_get(x_1249, 1); +lean_inc(x_1251); +lean_dec(x_1249); +x_1252 = lean_ctor_get(x_1243, 0); +lean_inc(x_1252); +x_1253 = lean_ctor_get(x_1243, 1); +lean_inc(x_1253); +lean_dec(x_1243); +x_1254 = lean_ctor_get(x_1250, 2); +lean_inc(x_1254); +lean_dec(x_1250); +x_1255 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1255, 0, x_1252); +lean_ctor_set(x_1255, 1, x_1253); +lean_ctor_set(x_1255, 2, x_1254); +x_1256 = lean_st_ref_get(x_7, x_1251); +lean_dec(x_7); +x_1257 = lean_ctor_get(x_1256, 1); +lean_inc(x_1257); +lean_dec(x_1256); +x_1258 = lean_st_ref_set(x_5, x_1255, x_1257); +lean_dec(x_5); +x_1259 = !lean_is_exclusive(x_1258); +if (x_1259 == 0) +{ +lean_object* x_1260; +x_1260 = lean_ctor_get(x_1258, 0); +lean_dec(x_1260); +lean_ctor_set(x_1258, 0, x_1245); +return x_1258; } 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; -x_168 = lean_ctor_get(x_164, 1); -lean_inc(x_168); -lean_dec(x_164); -x_169 = lean_ctor_get(x_165, 0); -lean_inc(x_169); -lean_dec(x_165); -x_170 = lean_alloc_closure((void*)(l_Lean_Compiler_visitMatch___boxed), 6, 2); -lean_closure_set(x_170, 0, x_9); -lean_closure_set(x_170, 1, x_169); -x_171 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___lambda__11), 8, 3); -lean_closure_set(x_171, 0, x_2); -lean_closure_set(x_171, 1, x_1); -lean_closure_set(x_171, 2, x_3); -x_172 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_172, 0, x_170); -lean_closure_set(x_172, 1, x_171); -x_173 = l_Lean_Compiler_withNewScopeImp___rarg(x_172, x_5, x_6, x_7, x_168); -return x_173; +lean_object* x_1261; lean_object* x_1262; +x_1261 = lean_ctor_get(x_1258, 1); +lean_inc(x_1261); +lean_dec(x_1258); +x_1262 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1262, 0, x_1245); +lean_ctor_set(x_1262, 1, x_1261); +return x_1262; } } -else +block_1282: { -uint8_t x_174; +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; uint8_t x_1278; +x_1266 = lean_st_ref_get(x_7, x_1265); +x_1267 = lean_ctor_get(x_1266, 1); +lean_inc(x_1267); +lean_dec(x_1266); +x_1268 = lean_st_ref_get(x_5, x_1267); +x_1269 = lean_ctor_get(x_1268, 0); +lean_inc(x_1269); +x_1270 = lean_ctor_get(x_1268, 1); +lean_inc(x_1270); +lean_dec(x_1268); +x_1271 = lean_ctor_get(x_1243, 0); +lean_inc(x_1271); +x_1272 = lean_ctor_get(x_1243, 1); +lean_inc(x_1272); +lean_dec(x_1243); +x_1273 = lean_ctor_get(x_1269, 2); +lean_inc(x_1273); +lean_dec(x_1269); +x_1274 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1274, 0, x_1271); +lean_ctor_set(x_1274, 1, x_1272); +lean_ctor_set(x_1274, 2, x_1273); +x_1275 = lean_st_ref_get(x_7, x_1270); +lean_dec(x_7); +x_1276 = lean_ctor_get(x_1275, 1); +lean_inc(x_1276); +lean_dec(x_1275); +x_1277 = lean_st_ref_set(x_5, x_1274, x_1276); +lean_dec(x_5); +x_1278 = !lean_is_exclusive(x_1277); +if (x_1278 == 0) +{ +lean_object* x_1279; +x_1279 = lean_ctor_get(x_1277, 0); +lean_dec(x_1279); +lean_ctor_set_tag(x_1277, 1); +lean_ctor_set(x_1277, 0, x_1264); +return x_1277; +} +else +{ +lean_object* x_1280; lean_object* x_1281; +x_1280 = lean_ctor_get(x_1277, 1); +lean_inc(x_1280); +lean_dec(x_1277); +x_1281 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1281, 0, x_1264); +lean_ctor_set(x_1281, 1, x_1280); +return x_1281; +} +} +} +} +else +{ +uint8_t x_1351; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -31446,23 +31149,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_174 = !lean_is_exclusive(x_164); -if (x_174 == 0) +x_1351 = !lean_is_exclusive(x_1234); +if (x_1351 == 0) { -return x_164; +return x_1234; } else { -lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_175 = lean_ctor_get(x_164, 0); -x_176 = lean_ctor_get(x_164, 1); -lean_inc(x_176); -lean_inc(x_175); -lean_dec(x_164); -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; +lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; +x_1352 = lean_ctor_get(x_1234, 0); +x_1353 = lean_ctor_get(x_1234, 1); +lean_inc(x_1353); +lean_inc(x_1352); +lean_dec(x_1234); +x_1354 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1354, 0, x_1352); +lean_ctor_set(x_1354, 1, x_1353); +return x_1354; } } } @@ -31470,51 +31173,270 @@ return x_177; } case 6: { -lean_object* x_178; +lean_object* x_1355; lean_dec(x_3); lean_dec(x_1); -x_178 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_8); -return x_178; +x_1355 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_8); +return x_1355; } case 8: { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +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; +x_1356 = lean_st_ref_get(x_7, x_8); +x_1357 = lean_ctor_get(x_1356, 1); +lean_inc(x_1357); +lean_dec(x_1356); +x_1358 = lean_st_ref_get(x_5, x_1357); +x_1359 = lean_ctor_get(x_1358, 0); +lean_inc(x_1359); +x_1360 = lean_ctor_get(x_1358, 1); +lean_inc(x_1360); +lean_dec(x_1358); +x_1361 = lean_st_ref_get(x_7, x_1360); +x_1362 = lean_ctor_get(x_1361, 1); +lean_inc(x_1362); +lean_dec(x_1361); +x_1363 = lean_st_ref_take(x_5, x_1362); +x_1364 = lean_ctor_get(x_1363, 0); +lean_inc(x_1364); +x_1365 = lean_ctor_get(x_1363, 1); +lean_inc(x_1365); +lean_dec(x_1363); +x_1366 = lean_ctor_get(x_1364, 0); +lean_inc(x_1366); +x_1367 = lean_ctor_get(x_1364, 2); +lean_inc(x_1367); +lean_dec(x_1364); +x_1368 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_1369 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1369, 0, x_1366); +lean_ctor_set(x_1369, 1, x_1368); +lean_ctor_set(x_1369, 2, x_1367); +x_1370 = lean_st_ref_set(x_5, x_1369, x_1365); +x_1371 = lean_ctor_get(x_1370, 1); +lean_inc(x_1371); +lean_dec(x_1370); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); lean_inc(x_3); -x_179 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLetImp), 6, 2); -lean_closure_set(x_179, 0, x_9); -lean_closure_set(x_179, 1, x_3); -x_180 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2___boxed), 8, 3); -lean_closure_set(x_180, 0, x_1); -lean_closure_set(x_180, 1, x_2); -lean_closure_set(x_180, 2, x_3); -x_181 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_181, 0, x_179); -lean_closure_set(x_181, 1, x_180); -x_182 = l_Lean_Compiler_withNewScopeImp___rarg(x_181, x_5, x_6, x_7, x_8); -return x_182; +x_1372 = l_Lean_Compiler_visitLetImp_go(x_3, x_9, x_1368, x_5, x_6, x_7, x_1371); +if (lean_obj_tag(x_1372) == 0) +{ +lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; +x_1373 = lean_ctor_get(x_1372, 0); +lean_inc(x_1373); +x_1374 = lean_ctor_get(x_1372, 1); +lean_inc(x_1374); +lean_dec(x_1372); +lean_inc(x_7); +lean_inc(x_5); +x_1375 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_1373, x_5, x_6, x_7, x_1374); +lean_dec(x_1373); +if (lean_obj_tag(x_1375) == 0) +{ +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; uint8_t x_1390; +x_1376 = lean_ctor_get(x_1375, 0); +lean_inc(x_1376); +x_1377 = lean_ctor_get(x_1375, 1); +lean_inc(x_1377); +lean_dec(x_1375); +x_1378 = lean_st_ref_get(x_7, x_1377); +x_1379 = lean_ctor_get(x_1378, 1); +lean_inc(x_1379); +lean_dec(x_1378); +x_1380 = lean_st_ref_get(x_5, x_1379); +x_1381 = lean_ctor_get(x_1380, 0); +lean_inc(x_1381); +x_1382 = lean_ctor_get(x_1380, 1); +lean_inc(x_1382); +lean_dec(x_1380); +x_1383 = lean_ctor_get(x_1359, 0); +lean_inc(x_1383); +x_1384 = lean_ctor_get(x_1359, 1); +lean_inc(x_1384); +lean_dec(x_1359); +x_1385 = lean_ctor_get(x_1381, 2); +lean_inc(x_1385); +lean_dec(x_1381); +x_1386 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1386, 0, x_1383); +lean_ctor_set(x_1386, 1, x_1384); +lean_ctor_set(x_1386, 2, x_1385); +x_1387 = lean_st_ref_get(x_7, x_1382); +lean_dec(x_7); +x_1388 = lean_ctor_get(x_1387, 1); +lean_inc(x_1388); +lean_dec(x_1387); +x_1389 = lean_st_ref_set(x_5, x_1386, x_1388); +lean_dec(x_5); +x_1390 = !lean_is_exclusive(x_1389); +if (x_1390 == 0) +{ +lean_object* x_1391; +x_1391 = lean_ctor_get(x_1389, 0); +lean_dec(x_1391); +lean_ctor_set(x_1389, 0, x_1376); +return x_1389; +} +else +{ +lean_object* x_1392; lean_object* x_1393; +x_1392 = lean_ctor_get(x_1389, 1); +lean_inc(x_1392); +lean_dec(x_1389); +x_1393 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1393, 0, x_1376); +lean_ctor_set(x_1393, 1, x_1392); +return x_1393; +} +} +else +{ +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; uint8_t x_1408; +x_1394 = lean_ctor_get(x_1375, 0); +lean_inc(x_1394); +x_1395 = lean_ctor_get(x_1375, 1); +lean_inc(x_1395); +lean_dec(x_1375); +x_1396 = lean_st_ref_get(x_7, x_1395); +x_1397 = lean_ctor_get(x_1396, 1); +lean_inc(x_1397); +lean_dec(x_1396); +x_1398 = lean_st_ref_get(x_5, x_1397); +x_1399 = lean_ctor_get(x_1398, 0); +lean_inc(x_1399); +x_1400 = lean_ctor_get(x_1398, 1); +lean_inc(x_1400); +lean_dec(x_1398); +x_1401 = lean_ctor_get(x_1359, 0); +lean_inc(x_1401); +x_1402 = lean_ctor_get(x_1359, 1); +lean_inc(x_1402); +lean_dec(x_1359); +x_1403 = lean_ctor_get(x_1399, 2); +lean_inc(x_1403); +lean_dec(x_1399); +x_1404 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1404, 0, x_1401); +lean_ctor_set(x_1404, 1, x_1402); +lean_ctor_set(x_1404, 2, x_1403); +x_1405 = lean_st_ref_get(x_7, x_1400); +lean_dec(x_7); +x_1406 = lean_ctor_get(x_1405, 1); +lean_inc(x_1406); +lean_dec(x_1405); +x_1407 = lean_st_ref_set(x_5, x_1404, x_1406); +lean_dec(x_5); +x_1408 = !lean_is_exclusive(x_1407); +if (x_1408 == 0) +{ +lean_object* x_1409; +x_1409 = lean_ctor_get(x_1407, 0); +lean_dec(x_1409); +lean_ctor_set_tag(x_1407, 1); +lean_ctor_set(x_1407, 0, x_1394); +return x_1407; +} +else +{ +lean_object* x_1410; lean_object* x_1411; +x_1410 = lean_ctor_get(x_1407, 1); +lean_inc(x_1410); +lean_dec(x_1407); +x_1411 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1411, 0, x_1394); +lean_ctor_set(x_1411, 1, x_1410); +return x_1411; +} +} +} +else +{ +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; uint8_t x_1426; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1412 = lean_ctor_get(x_1372, 0); +lean_inc(x_1412); +x_1413 = lean_ctor_get(x_1372, 1); +lean_inc(x_1413); +lean_dec(x_1372); +x_1414 = lean_st_ref_get(x_7, x_1413); +x_1415 = lean_ctor_get(x_1414, 1); +lean_inc(x_1415); +lean_dec(x_1414); +x_1416 = lean_st_ref_get(x_5, x_1415); +x_1417 = lean_ctor_get(x_1416, 0); +lean_inc(x_1417); +x_1418 = lean_ctor_get(x_1416, 1); +lean_inc(x_1418); +lean_dec(x_1416); +x_1419 = lean_ctor_get(x_1359, 0); +lean_inc(x_1419); +x_1420 = lean_ctor_get(x_1359, 1); +lean_inc(x_1420); +lean_dec(x_1359); +x_1421 = lean_ctor_get(x_1417, 2); +lean_inc(x_1421); +lean_dec(x_1417); +x_1422 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1422, 0, x_1419); +lean_ctor_set(x_1422, 1, x_1420); +lean_ctor_set(x_1422, 2, x_1421); +x_1423 = lean_st_ref_get(x_7, x_1418); +lean_dec(x_7); +x_1424 = lean_ctor_get(x_1423, 1); +lean_inc(x_1424); +lean_dec(x_1423); +x_1425 = lean_st_ref_set(x_5, x_1422, x_1424); +lean_dec(x_5); +x_1426 = !lean_is_exclusive(x_1425); +if (x_1426 == 0) +{ +lean_object* x_1427; +x_1427 = lean_ctor_get(x_1425, 0); +lean_dec(x_1427); +lean_ctor_set_tag(x_1425, 1); +lean_ctor_set(x_1425, 0, x_1412); +return x_1425; +} +else +{ +lean_object* x_1428; lean_object* x_1429; +x_1428 = lean_ctor_get(x_1425, 1); +lean_inc(x_1428); +lean_dec(x_1425); +x_1429 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1429, 0, x_1412); +lean_ctor_set(x_1429, 1, x_1428); +return x_1429; +} +} } case 10: { -lean_object* x_183; lean_object* x_184; +lean_object* x_1430; lean_object* x_1431; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_183 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; -x_184 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_183, x_5, x_6, x_7, x_8); -return x_184; +x_1430 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2; +x_1431 = l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1(x_1430, x_5, x_6, x_7, x_8); +return x_1431; } case 11: { -lean_object* x_185; +lean_object* x_1432; lean_dec(x_3); lean_dec(x_1); -x_185 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_8); -return x_185; +x_1432 = lean_apply_5(x_2, x_9, x_5, x_6, x_7, x_8); +return x_1432; } default: { -lean_object* x_186; lean_object* x_187; +lean_object* x_1433; lean_object* x_1434; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -31522,39 +31444,431 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_186 = lean_box(0); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_8); -return x_187; +x_1433 = lean_box(0); +x_1434 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1434, 0, x_1433); +lean_ctor_set(x_1434, 1, x_8); +return x_1434; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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; lean_object* x_10; -x_9 = lean_ctor_get(x_4, 1); -x_10 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); -return x_10; +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, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_st_ref_get(x_5, 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_st_ref_get(x_7, x_13); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_take(x_5, x_15); +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_is_exclusive(x_17); +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 = lean_ctor_get(x_17, 1); +lean_dec(x_20); +x_21 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +lean_ctor_set(x_17, 1, x_21); +x_22 = lean_st_ref_set(x_5, x_17, x_18); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Compiler_visitLambda(x_4, x_5, x_6, x_7, x_23); +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_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +lean_inc(x_7); +lean_inc(x_5); +x_28 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_27, x_5, x_6, x_7, x_26); +lean_dec(x_27); +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; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +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_get(x_7, x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_st_ref_get(x_5, 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); +lean_dec(x_33); +x_36 = lean_ctor_get(x_12, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_dec(x_12); +x_38 = !lean_is_exclusive(x_34); +if (x_38 == 0) +{ +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_39 = lean_ctor_get(x_34, 1); +lean_dec(x_39); +x_40 = lean_ctor_get(x_34, 0); +lean_dec(x_40); +lean_ctor_set(x_34, 1, x_37); +lean_ctor_set(x_34, 0, x_36); +x_41 = lean_st_ref_get(x_7, x_35); +lean_dec(x_7); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_st_ref_set(x_5, x_34, x_42); +lean_dec(x_5); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +lean_ctor_set(x_43, 0, x_29); +return x_43; } +else +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_29); +lean_ctor_set(x_47, 1, x_46); +return x_47; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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: +} +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_9, 0, x_4); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__1___lambda__1___boxed), 8, 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); -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Compiler_withNewScopeImp___rarg(x_11, x_5, x_6, x_7, x_8); -return x_12; +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_ctor_get(x_34, 2); +lean_inc(x_48); +lean_dec(x_34); +x_49 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_37); +lean_ctor_set(x_49, 2, x_48); +x_50 = lean_st_ref_get(x_7, x_35); +lean_dec(x_7); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = lean_st_ref_set(x_5, x_49, x_51); +lean_dec(x_5); +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_29); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +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; uint8_t x_65; +x_56 = lean_ctor_get(x_28, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_28, 1); +lean_inc(x_57); +lean_dec(x_28); +x_58 = lean_st_ref_get(x_7, x_57); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_st_ref_get(x_5, x_59); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_ctor_get(x_12, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_12, 1); +lean_inc(x_64); +lean_dec(x_12); +x_65 = !lean_is_exclusive(x_61); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_66 = lean_ctor_get(x_61, 1); +lean_dec(x_66); +x_67 = lean_ctor_get(x_61, 0); +lean_dec(x_67); +lean_ctor_set(x_61, 1, x_64); +lean_ctor_set(x_61, 0, x_63); +x_68 = lean_st_ref_get(x_7, x_62); +lean_dec(x_7); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = lean_st_ref_set(x_5, x_61, x_69); +lean_dec(x_5); +x_71 = !lean_is_exclusive(x_70); +if (x_71 == 0) +{ +lean_object* x_72; +x_72 = lean_ctor_get(x_70, 0); +lean_dec(x_72); +lean_ctor_set_tag(x_70, 1); +lean_ctor_set(x_70, 0, x_56); +return x_70; +} +else +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_70, 1); +lean_inc(x_73); +lean_dec(x_70); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_56); +lean_ctor_set(x_74, 1, x_73); +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; +x_75 = lean_ctor_get(x_61, 2); +lean_inc(x_75); +lean_dec(x_61); +x_76 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_76, 0, x_63); +lean_ctor_set(x_76, 1, x_64); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_st_ref_get(x_7, x_62); +lean_dec(x_7); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = lean_st_ref_set(x_5, x_76, x_78); +lean_dec(x_5); +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_81 = x_79; +} else { + lean_dec_ref(x_79); + x_81 = lean_box(0); +} +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(1, 2, 0); +} else { + x_82 = x_81; + lean_ctor_set_tag(x_82, 1); +} +lean_ctor_set(x_82, 0, x_56); +lean_ctor_set(x_82, 1, x_80); +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_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_83 = lean_ctor_get(x_17, 0); +x_84 = lean_ctor_get(x_17, 2); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_17); +x_85 = l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1; +x_86 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = lean_st_ref_set(x_5, x_86, x_18); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = l_Lean_Compiler_visitLambda(x_4, x_5, x_6, x_7, x_88); +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_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +lean_inc(x_7); +lean_inc(x_5); +x_93 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__2(x_1, x_2, x_3, x_92, x_5, x_6, x_7, x_91); +lean_dec(x_92); +if (lean_obj_tag(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; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_96 = lean_st_ref_get(x_7, x_95); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = lean_st_ref_get(x_5, x_97); +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_101 = lean_ctor_get(x_12, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_12, 1); +lean_inc(x_102); +lean_dec(x_12); +x_103 = lean_ctor_get(x_99, 2); +lean_inc(x_103); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + lean_ctor_release(x_99, 2); + x_104 = x_99; +} else { + lean_dec_ref(x_99); + x_104 = lean_box(0); +} +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(0, 3, 0); +} else { + x_105 = x_104; +} +lean_ctor_set(x_105, 0, x_101); +lean_ctor_set(x_105, 1, x_102); +lean_ctor_set(x_105, 2, x_103); +x_106 = lean_st_ref_get(x_7, x_100); +lean_dec(x_7); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +lean_dec(x_106); +x_108 = lean_st_ref_set(x_5, x_105, x_107); +lean_dec(x_5); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_110 = x_108; +} else { + lean_dec_ref(x_108); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_94); +lean_ctor_set(x_111, 1, x_109); +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_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; +x_112 = lean_ctor_get(x_93, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_93, 1); +lean_inc(x_113); +lean_dec(x_93); +x_114 = lean_st_ref_get(x_7, x_113); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_get(x_5, x_115); +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_12, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_12, 1); +lean_inc(x_120); +lean_dec(x_12); +x_121 = lean_ctor_get(x_117, 2); +lean_inc(x_121); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + 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, 3, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_119); +lean_ctor_set(x_123, 1, x_120); +lean_ctor_set(x_123, 2, x_121); +x_124 = lean_st_ref_get(x_7, x_118); +lean_dec(x_7); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_126 = lean_st_ref_set(x_5, x_123, x_125); +lean_dec(x_5); +x_127 = lean_ctor_get(x_126, 1); +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); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); +} else { + x_129 = x_128; + lean_ctor_set_tag(x_129, 1); +} +lean_ctor_set(x_129, 0, x_112); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +} } } LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { @@ -31652,54 +31966,55 @@ return x_17; } else { -lean_object* x_18; +lean_object* x_18; lean_object* x_19; +x_18 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1; lean_inc(x_1); -x_18 = l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp(x_1, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_18) == 0) +x_19 = l_Lean_Compiler_JoinPoints_forEachFVar___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3(x_1, x_18, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_19) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_1); -return x_18; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_1); +return x_19; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_1); -lean_ctor_set(x_22, 1, x_21); -return x_22; +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_1); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } else { -uint8_t x_23; +uint8_t x_24; lean_dec(x_1); -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -return x_18; +return x_19; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); +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_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -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_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; } } } @@ -32294,15 +32609,6 @@ lean_dec(x_4); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___at_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_4); -return x_9; -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goLetValue___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: { @@ -32348,6 +32654,8 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_CompilerM(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1 = _init_l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___lambda__3___closed__1); l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__1 = _init_l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__1(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__1); l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__2 = _init_l_Lean_Compiler_JoinPoints_forEachFVar___rarg___closed__2(); @@ -32362,46 +32670,64 @@ l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___ra lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__1); l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2 = _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2(); lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitTails___rarg___closed__2); +l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1 = _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_visitLambda___rarg___closed__1); +l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__1); +l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo___closed__2); +l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_instInhabitedCandidateInfo); l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1(); lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__1); l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2(); lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__2); l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3(); lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__3); +l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4(); +lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__2___closed__4); +l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__1); +l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2 = _init_l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___spec__4___closed__2); l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1 = _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__1); l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2 = _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2(); lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___lambda__1___closed__2); l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1 = _init_l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_JoinPoints_0__Lean_Compiler_JoinPoints_JoinPointFinder_removeCandidatesContainedIn___closed__1); +l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__1); +l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2 = _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__2); +l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__3 = _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__3(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__3); +l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4 = _init_l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__2___closed__4); l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__1); l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___closed__2); -l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__1); -l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__2); -l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___lambda__2___closed__3); -l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1 = _init_l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1(); -lean_mark_persistent(l_Std_HashMap_toArray___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints___spec__3___closed__1); -l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1(); -lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1); -l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1 = _init_l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__3___closed__1); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__1 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__1(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__1); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__2); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__3); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__4); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__5); -l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6(); -lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__4___closed__6); +l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__1); +l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__2); +l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3 = _init_l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3(); +lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___closed__3); +l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__1___closed__1); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__1); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__2); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__3); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__4); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__5); +l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6 = _init_l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2___closed__6); l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__1); l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__2(); @@ -32410,20 +32736,8 @@ l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3 lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__3); l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__1___closed__4); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___lambda__2___closed__1); l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__1); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__2); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__3); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__4); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__5); -l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6(); -lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___closed__6); l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1(); lean_mark_persistent(l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__1); l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2 = _init_l_Lean_Compiler_JoinPoints_JoinPointChecker_checkJoinPoints_goTailApp___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF.c b/stage0/stdlib/Lean/Compiler/LCNF.c index 332c4fbacfcc..0f5297d6fc9d 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF.c +++ b/stage0/stdlib/Lean/Compiler/LCNF.c @@ -15,12 +15,12 @@ extern "C" { #endif static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4___closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_mkFreshFVarId___at_Lean_Compiler_ToLCNF_mkLocalDecl___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitLet(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_typeCache___default; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__15; lean_object* lean_erase_macro_scopes(lean_object*); @@ -50,14 +50,12 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_lctx___default; static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4___closed__2; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__5; static lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__1___closed__2; -static size_t l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__1; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_mk_let_decl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static size_t l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__4; uint8_t l_Lean_Compiler_compatibleTypes(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitFalseRec___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -91,26 +89,34 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ToLCNF_toLCNF static lean_object* l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_etaIfUnderApplied___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(lean_object*, lean_object*); +static lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1; static lean_object* l_Lean_Compiler_ToLCNF_mkLetUsingScope___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAppArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__10; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_mkLcCast___at_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___spec__1___closed__2; static lean_object* l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__1; static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkFreshBinderName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_visitBoundedLambda_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAppArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_etaExpandN___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_Compiler_ToLCNF_toLCNF_visitAlt___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_Compiler_ToLCNF_toLCNF_visitLambda___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_inferType___boxed(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_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_State_lctx___default___closed__6; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__6; @@ -120,8 +126,10 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_liftMetaM(lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitLambda(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3___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_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Compiler_ToLCNF_etaExpandN___spec__1(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -142,6 +150,7 @@ static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___closed__1; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__3; static lean_object* l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_visitLambda_go(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__24; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_nextIdx___default; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -169,7 +178,7 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toL lean_object* l_Lean_Compiler_toLCNFType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitLambda___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__7; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__5; lean_object* l_Lean_Compiler_getCasesInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -203,6 +212,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1 LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_letFVars___default; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__1___closed__1; +static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visit___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -212,6 +222,7 @@ static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__1; extern lean_object* l_Lean_Expr_instHashableExpr; uint64_t l_Lean_Expr_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_State_lctx___default___closed__7; lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -219,6 +230,7 @@ static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__11; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ToLCNF_toLCNF_visitApp___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__1___closed__2; lean_object* l_Lean_Compiler_mkLcProof(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_ToLCNF_State_cache___default___lambda__1(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__1___boxed(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_Lean_Compiler_ToLCNF_toLCNF_visitCases___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,7 +261,6 @@ lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_ToLCNF_toLCNF_visitLambda___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeFormerType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitApp___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAndRec___boxed(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_getConstInfo___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -264,6 +275,8 @@ size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_projectionFnInfoExt; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAlt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNFType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAppDefault___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_mk_local_decl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); @@ -286,7 +299,7 @@ static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__12; static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_withNewRootScope___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); lean_object* l_Lean_Expr_beta(lean_object*, lean_object*); @@ -302,6 +315,7 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_ToLCNF_mkLocalDecl_ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___closed__2; extern lean_object* l_Lean_Core_instMonadCoreM; +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Compiler_ToLCNF_mkLocalDecl___spec__2(uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___spec__1___closed__3; @@ -334,7 +348,9 @@ static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitFalseRec___closed__1; static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__4; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_visitLambda_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instBEq___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_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visit___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__8; @@ -345,9 +361,11 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLetDecl___boxed(lean_object*, lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__21; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_etaExpandN(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__5; extern lean_object* l_Lean_Expr_instBEqExpr; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Compiler_ToLCNF_etaExpandN___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -363,18 +381,20 @@ lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAlt___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Compiler_ToLCNF_mkLocalDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__23; -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_State_lctx___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLetUsingScope___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLetUsingScope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ToLCNF_toLCNF_visitApp___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_ToLCNF_toLCNF___closed__2; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNFType(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAppArg___lambda__1(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* l_instHashableBool___boxed(lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__4; @@ -388,6 +408,8 @@ lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__9; static lean_object* l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__27; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitProj___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_visitBoundedLambda(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -396,9 +418,11 @@ extern lean_object* l_Lean_Compiler_anyTypeExpr; lean_object* l_Lean_Expr_isConstructorApp_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__11; static lean_object* l_Lean_Compiler_ToLCNF_State_lctx___default___closed__4; +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__5; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitAlt___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitChild___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__17; @@ -417,9 +441,11 @@ lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_withRoot___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitMData___boxed(lean_object*); +static size_t l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Compiler_ToLCNF_toLCNF_visitApp___spec__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__28; +static lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___closed__5; static lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__7; static lean_object* _init_l_Lean_Compiler_ToLCNF_State_lctx___default___closed__1() { _start: @@ -540,9 +566,83 @@ x_1 = lean_unsigned_to_nat(1u); return x_1; } } +LEAN_EXPORT uint8_t l_Lean_Compiler_ToLCNF_State_cache___default___lambda__1(uint8_t x_1, uint8_t x_2) { +_start: +{ +if (x_1 == 0) +{ +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; +} +} +else +{ +return x_2; +} +} +} static lean_object* _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_State_cache___default___lambda__1___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Expr_instBEqExpr; +x_2 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__2; +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_Lean_Compiler_ToLCNF_State_cache___default___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instHashableBool___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Expr_instHashableExpr; +x_2 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__4; +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_Lean_Compiler_ToLCNF_State_cache___default___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_ToLCNF_State_lctx___default___closed__2; x_2 = lean_unsigned_to_nat(0u); @@ -556,7 +656,40 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_State_cache___default() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__1; +x_1 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__6; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_State_cache___default___lambda__1___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_Compiler_ToLCNF_State_cache___default___lambda__1(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_ToLCNF_State_lctx___default___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_Compiler_ToLCNF_State_typeCache___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1; return x_1; } } @@ -746,7 +879,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__12() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__6; x_2 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__8; -x_3 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__1; +x_3 = l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1; x_4 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__11; x_5 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_5, 0, x_1); @@ -969,7 +1102,7 @@ lean_inc(x_3); x_24 = lean_apply_5(x_1, x_23, x_3, x_4, x_5, 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; 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; +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; uint8_t x_42; x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_24, 1); @@ -985,323 +1118,294 @@ lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_ctor_get(x_10, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_10, 1); +x_32 = lean_st_ref_get(x_5, x_31); +x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_10, 2); -lean_inc(x_34); -x_35 = lean_ctor_get(x_10, 4); +lean_dec(x_32); +x_34 = lean_st_ref_get(x_3, x_33); +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_10, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_10, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_10, 2); +lean_inc(x_39); +x_40 = lean_ctor_get(x_10, 4); +lean_inc(x_40); lean_dec(x_10); -x_36 = !lean_is_exclusive(x_30); -if (x_36 == 0) +x_41 = lean_ctor_get(x_30, 3); +lean_inc(x_41); +lean_dec(x_30); +x_42 = !lean_is_exclusive(x_35); +if (x_42 == 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_30, 4); -lean_dec(x_37); -x_38 = lean_ctor_get(x_30, 2); -lean_dec(x_38); -x_39 = lean_ctor_get(x_30, 1); -lean_dec(x_39); -x_40 = lean_ctor_get(x_30, 0); -lean_dec(x_40); -lean_ctor_set(x_30, 4, x_35); -lean_ctor_set(x_30, 2, x_34); -lean_ctor_set(x_30, 1, x_33); -lean_ctor_set(x_30, 0, x_32); -x_41 = lean_st_ref_get(x_5, x_31); -lean_dec(x_5); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -x_43 = lean_st_ref_set(x_3, x_30, 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; +x_43 = lean_ctor_get(x_35, 4); +lean_dec(x_43); +x_44 = lean_ctor_get(x_35, 3); +lean_dec(x_44); +x_45 = lean_ctor_get(x_35, 2); +lean_dec(x_45); +x_46 = lean_ctor_get(x_35, 1); +lean_dec(x_46); +x_47 = lean_ctor_get(x_35, 0); +lean_dec(x_47); +lean_ctor_set(x_35, 4, x_40); +lean_ctor_set(x_35, 3, x_41); +lean_ctor_set(x_35, 2, x_39); +lean_ctor_set(x_35, 1, x_38); +lean_ctor_set(x_35, 0, x_37); +x_48 = lean_st_ref_get(x_5, x_36); +lean_dec(x_5); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = lean_st_ref_set(x_3, x_35, x_49); lean_dec(x_3); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_43, 0); -lean_dec(x_45); -lean_ctor_set(x_43, 0, x_25); -return x_43; +lean_object* x_52; +x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_25); +return x_50; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_25); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_50, 1); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } 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_ctor_get(x_30, 3); -lean_inc(x_48); -lean_dec(x_30); -x_49 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_49, 0, x_32); -lean_ctor_set(x_49, 1, x_33); -lean_ctor_set(x_49, 2, x_34); -lean_ctor_set(x_49, 3, x_48); -lean_ctor_set(x_49, 4, x_35); -x_50 = lean_st_ref_get(x_5, x_31); -lean_dec(x_5); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_st_ref_set(x_3, x_49, x_51); +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_35, 5); +lean_inc(x_55); +lean_dec(x_35); +x_56 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_56, 0, x_37); +lean_ctor_set(x_56, 1, x_38); +lean_ctor_set(x_56, 2, x_39); +lean_ctor_set(x_56, 3, x_41); +lean_ctor_set(x_56, 4, x_40); +lean_ctor_set(x_56, 5, x_55); +x_57 = lean_st_ref_get(x_5, x_36); +lean_dec(x_5); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_st_ref_set(x_3, x_56, x_58); lean_dec(x_3); -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; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_61 = x_59; } else { - lean_dec_ref(x_52); - x_54 = lean_box(0); + lean_dec_ref(x_59); + x_61 = lean_box(0); } -if (lean_is_scalar(x_54)) { - x_55 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_54; + x_62 = x_61; } -lean_ctor_set(x_55, 0, x_25); -lean_ctor_set(x_55, 1, x_53); -return x_55; +lean_ctor_set(x_62, 0, x_25); +lean_ctor_set(x_62, 1, x_60); +return x_62; } } 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_object* x_65; lean_object* x_66; uint8_t x_67; -x_56 = lean_ctor_get(x_24, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_24, 1); -lean_inc(x_57); -lean_dec(x_24); -x_58 = lean_st_ref_get(x_5, x_57); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_st_ref_get(x_3, x_59); -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_ctor_get(x_10, 0); +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; uint8_t x_80; +x_63 = lean_ctor_get(x_24, 0); lean_inc(x_63); -x_64 = lean_ctor_get(x_10, 1); +x_64 = lean_ctor_get(x_24, 1); lean_inc(x_64); -x_65 = lean_ctor_get(x_10, 2); -lean_inc(x_65); -x_66 = lean_ctor_get(x_10, 4); +lean_dec(x_24); +x_65 = lean_st_ref_get(x_5, x_64); +x_66 = lean_ctor_get(x_65, 1); lean_inc(x_66); -lean_dec(x_10); -x_67 = !lean_is_exclusive(x_61); -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; uint8_t x_75; -x_68 = lean_ctor_get(x_61, 4); -lean_dec(x_68); -x_69 = lean_ctor_get(x_61, 2); -lean_dec(x_69); -x_70 = lean_ctor_get(x_61, 1); +lean_dec(x_65); +x_67 = lean_st_ref_get(x_3, x_66); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_st_ref_get(x_5, x_69); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); lean_dec(x_70); -x_71 = lean_ctor_get(x_61, 0); -lean_dec(x_71); -lean_ctor_set(x_61, 4, x_66); -lean_ctor_set(x_61, 2, x_65); -lean_ctor_set(x_61, 1, x_64); -lean_ctor_set(x_61, 0, x_63); -x_72 = lean_st_ref_get(x_5, x_62); -lean_dec(x_5); -x_73 = lean_ctor_get(x_72, 1); +x_72 = lean_st_ref_get(x_3, x_71); +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_74 = lean_st_ref_set(x_3, x_61, x_73); +x_75 = lean_ctor_get(x_10, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_10, 1); +lean_inc(x_76); +x_77 = lean_ctor_get(x_10, 2); +lean_inc(x_77); +x_78 = lean_ctor_get(x_10, 4); +lean_inc(x_78); +lean_dec(x_10); +x_79 = lean_ctor_get(x_68, 3); +lean_inc(x_79); +lean_dec(x_68); +x_80 = !lean_is_exclusive(x_73); +if (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; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_81 = lean_ctor_get(x_73, 4); +lean_dec(x_81); +x_82 = lean_ctor_get(x_73, 3); +lean_dec(x_82); +x_83 = lean_ctor_get(x_73, 2); +lean_dec(x_83); +x_84 = lean_ctor_get(x_73, 1); +lean_dec(x_84); +x_85 = lean_ctor_get(x_73, 0); +lean_dec(x_85); +lean_ctor_set(x_73, 4, x_78); +lean_ctor_set(x_73, 3, x_79); +lean_ctor_set(x_73, 2, x_77); +lean_ctor_set(x_73, 1, x_76); +lean_ctor_set(x_73, 0, x_75); +x_86 = lean_st_ref_get(x_5, x_74); +lean_dec(x_5); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = lean_st_ref_set(x_3, x_73, x_87); lean_dec(x_3); -x_75 = !lean_is_exclusive(x_74); -if (x_75 == 0) +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) { -lean_object* x_76; -x_76 = lean_ctor_get(x_74, 0); -lean_dec(x_76); -lean_ctor_set_tag(x_74, 1); -lean_ctor_set(x_74, 0, x_56); -return x_74; +lean_object* x_90; +x_90 = lean_ctor_get(x_88, 0); +lean_dec(x_90); +lean_ctor_set_tag(x_88, 1); +lean_ctor_set(x_88, 0, x_63); +return x_88; } else { -lean_object* x_77; lean_object* x_78; -x_77 = lean_ctor_get(x_74, 1); -lean_inc(x_77); -lean_dec(x_74); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_56); -lean_ctor_set(x_78, 1, x_77); -return x_78; +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_88, 1); +lean_inc(x_91); +lean_dec(x_88); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_63); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } 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; -x_79 = lean_ctor_get(x_61, 3); -lean_inc(x_79); -lean_dec(x_61); -x_80 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_80, 0, x_63); -lean_ctor_set(x_80, 1, x_64); -lean_ctor_set(x_80, 2, x_65); -lean_ctor_set(x_80, 3, x_79); -lean_ctor_set(x_80, 4, x_66); -x_81 = lean_st_ref_get(x_5, x_62); -lean_dec(x_5); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -lean_dec(x_81); -x_83 = lean_st_ref_set(x_3, x_80, x_82); +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_93 = lean_ctor_get(x_73, 5); +lean_inc(x_93); +lean_dec(x_73); +x_94 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_94, 0, x_75); +lean_ctor_set(x_94, 1, x_76); +lean_ctor_set(x_94, 2, x_77); +lean_ctor_set(x_94, 3, x_79); +lean_ctor_set(x_94, 4, x_78); +lean_ctor_set(x_94, 5, x_93); +x_95 = lean_st_ref_get(x_5, x_74); +lean_dec(x_5); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = lean_st_ref_set(x_3, x_94, x_96); lean_dec(x_3); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_85 = x_83; +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_99 = x_97; } else { - lean_dec_ref(x_83); - x_85 = lean_box(0); + lean_dec_ref(x_97); + x_99 = lean_box(0); } -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_99)) { + x_100 = lean_alloc_ctor(1, 2, 0); } else { - x_86 = x_85; - lean_ctor_set_tag(x_86, 1); + x_100 = x_99; + lean_ctor_set_tag(x_100, 1); } -lean_ctor_set(x_86, 0, x_56); -lean_ctor_set(x_86, 1, x_84); -return x_86; +lean_ctor_set(x_100, 0, x_63); +lean_ctor_set(x_100, 1, x_98); +return x_100; } } } 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; uint8_t x_95; lean_object* x_96; lean_object* x_97; -x_87 = lean_ctor_get(x_15, 0); -x_88 = lean_ctor_get(x_15, 1); -x_89 = lean_ctor_get(x_15, 3); -x_90 = lean_ctor_get(x_15, 4); -lean_inc(x_90); -lean_inc(x_89); -lean_inc(x_88); -lean_inc(x_87); +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; uint8_t x_110; lean_object* x_111; lean_object* x_112; +x_101 = lean_ctor_get(x_15, 0); +x_102 = lean_ctor_get(x_15, 1); +x_103 = lean_ctor_get(x_15, 3); +x_104 = lean_ctor_get(x_15, 4); +x_105 = lean_ctor_get(x_15, 5); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); +lean_inc(x_101); lean_dec(x_15); -x_91 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; -x_92 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_92, 0, x_87); -lean_ctor_set(x_92, 1, x_88); -lean_ctor_set(x_92, 2, x_91); -lean_ctor_set(x_92, 3, x_89); -lean_ctor_set(x_92, 4, x_90); -x_93 = lean_st_ref_set(x_3, x_92, x_16); -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec(x_93); -x_95 = 1; -x_96 = lean_box(x_95); +x_106 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +x_107 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_107, 0, x_101); +lean_ctor_set(x_107, 1, x_102); +lean_ctor_set(x_107, 2, x_106); +lean_ctor_set(x_107, 3, x_103); +lean_ctor_set(x_107, 4, x_104); +lean_ctor_set(x_107, 5, x_105); +x_108 = lean_st_ref_set(x_3, x_107, x_16); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec(x_108); +x_110 = 1; +x_111 = lean_box(x_110); lean_inc(x_5); lean_inc(x_3); -x_97 = lean_apply_5(x_1, x_96, x_3, x_4, x_5, x_94); -if (lean_obj_tag(x_97) == 0) +x_112 = lean_apply_5(x_1, x_111, x_3, x_4, x_5, x_109); +if (lean_obj_tag(x_112) == 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; 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_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = lean_st_ref_get(x_5, x_99); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -lean_dec(x_100); -x_102 = lean_st_ref_get(x_3, x_101); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_105 = lean_ctor_get(x_10, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_10, 1); -lean_inc(x_106); -x_107 = lean_ctor_get(x_10, 2); -lean_inc(x_107); -x_108 = lean_ctor_get(x_10, 4); -lean_inc(x_108); -lean_dec(x_10); -x_109 = lean_ctor_get(x_103, 3); -lean_inc(x_109); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - lean_ctor_release(x_103, 2); - lean_ctor_release(x_103, 3); - lean_ctor_release(x_103, 4); - x_110 = x_103; -} else { - lean_dec_ref(x_103); - x_110 = lean_box(0); -} -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(0, 5, 0); -} else { - x_111 = x_110; -} -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_106); -lean_ctor_set(x_111, 2, x_107); -lean_ctor_set(x_111, 3, x_109); -lean_ctor_set(x_111, 4, x_108); -x_112 = lean_st_ref_get(x_5, x_104); -lean_dec(x_5); -x_113 = lean_ctor_get(x_112, 1); +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; +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_114 = lean_st_ref_set(x_3, x_111, x_113); -lean_dec(x_3); -x_115 = lean_ctor_get(x_114, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_114)) { - lean_ctor_release(x_114, 0); - lean_ctor_release(x_114, 1); - x_116 = x_114; -} else { - lean_dec_ref(x_114); - x_116 = lean_box(0); -} -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(0, 2, 0); -} else { - x_117 = x_116; -} -lean_ctor_set(x_117, 0, x_98); -lean_ctor_set(x_117, 1, x_115); -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_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; -x_118 = lean_ctor_get(x_97, 0); +x_115 = lean_st_ref_get(x_5, x_114); +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +x_117 = lean_st_ref_get(x_3, x_116); +x_118 = lean_ctor_get(x_117, 0); lean_inc(x_118); -x_119 = lean_ctor_get(x_97, 1); +x_119 = lean_ctor_get(x_117, 1); lean_inc(x_119); -lean_dec(x_97); +lean_dec(x_117); x_120 = lean_st_ref_get(x_5, x_119); x_121 = lean_ctor_get(x_120, 1); lean_inc(x_121); @@ -1321,74 +1425,1314 @@ lean_inc(x_127); x_128 = lean_ctor_get(x_10, 4); lean_inc(x_128); lean_dec(x_10); -x_129 = lean_ctor_get(x_123, 3); +x_129 = lean_ctor_get(x_118, 3); lean_inc(x_129); +lean_dec(x_118); +x_130 = lean_ctor_get(x_123, 5); +lean_inc(x_130); 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); lean_ctor_release(x_123, 4); - x_130 = x_123; + lean_ctor_release(x_123, 5); + x_131 = x_123; } else { lean_dec_ref(x_123); - x_130 = lean_box(0); + x_131 = lean_box(0); } -if (lean_is_scalar(x_130)) { - x_131 = lean_alloc_ctor(0, 5, 0); +if (lean_is_scalar(x_131)) { + x_132 = lean_alloc_ctor(0, 6, 0); } else { - x_131 = x_130; + x_132 = x_131; } -lean_ctor_set(x_131, 0, x_125); -lean_ctor_set(x_131, 1, x_126); -lean_ctor_set(x_131, 2, x_127); -lean_ctor_set(x_131, 3, x_129); -lean_ctor_set(x_131, 4, x_128); -x_132 = lean_st_ref_get(x_5, x_124); +lean_ctor_set(x_132, 0, x_125); +lean_ctor_set(x_132, 1, x_126); +lean_ctor_set(x_132, 2, x_127); +lean_ctor_set(x_132, 3, x_129); +lean_ctor_set(x_132, 4, x_128); +lean_ctor_set(x_132, 5, x_130); +x_133 = lean_st_ref_get(x_5, x_124); lean_dec(x_5); -x_133 = lean_ctor_get(x_132, 1); -lean_inc(x_133); -lean_dec(x_132); -x_134 = lean_st_ref_set(x_3, x_131, x_133); +x_134 = lean_ctor_get(x_133, 1); +lean_inc(x_134); +lean_dec(x_133); +x_135 = lean_st_ref_set(x_3, x_132, x_134); lean_dec(x_3); -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_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; +} else { + lean_dec_ref(x_135); + x_137 = lean_box(0); +} +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_137; +} +lean_ctor_set(x_138, 0, x_113); +lean_ctor_set(x_138, 1, x_136); +return x_138; +} +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; 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_139 = lean_ctor_get(x_112, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_112, 1); +lean_inc(x_140); +lean_dec(x_112); +x_141 = lean_st_ref_get(x_5, x_140); +x_142 = lean_ctor_get(x_141, 1); +lean_inc(x_142); +lean_dec(x_141); +x_143 = lean_st_ref_get(x_3, x_142); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = lean_st_ref_get(x_5, x_145); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +lean_dec(x_146); +x_148 = lean_st_ref_get(x_3, x_147); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = lean_ctor_get(x_10, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_10, 1); +lean_inc(x_152); +x_153 = lean_ctor_get(x_10, 2); +lean_inc(x_153); +x_154 = lean_ctor_get(x_10, 4); +lean_inc(x_154); +lean_dec(x_10); +x_155 = lean_ctor_get(x_144, 3); +lean_inc(x_155); +lean_dec(x_144); +x_156 = lean_ctor_get(x_149, 5); +lean_inc(x_156); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + lean_ctor_release(x_149, 2); + lean_ctor_release(x_149, 3); + lean_ctor_release(x_149, 4); + lean_ctor_release(x_149, 5); + x_157 = x_149; +} else { + lean_dec_ref(x_149); + x_157 = lean_box(0); +} +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(0, 6, 0); +} else { + x_158 = x_157; +} +lean_ctor_set(x_158, 0, x_151); +lean_ctor_set(x_158, 1, x_152); +lean_ctor_set(x_158, 2, x_153); +lean_ctor_set(x_158, 3, x_155); +lean_ctor_set(x_158, 4, x_154); +lean_ctor_set(x_158, 5, x_156); +x_159 = lean_st_ref_get(x_5, x_150); +lean_dec(x_5); +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_160); +lean_dec(x_159); +x_161 = lean_st_ref_set(x_3, x_158, x_160); +lean_dec(x_3); +x_162 = lean_ctor_get(x_161, 1); +lean_inc(x_162); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_163 = x_161; } else { - lean_dec_ref(x_134); - x_136 = lean_box(0); + lean_dec_ref(x_161); + x_163 = lean_box(0); } -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_163)) { + x_164 = lean_alloc_ctor(1, 2, 0); } else { - x_137 = x_136; - lean_ctor_set_tag(x_137, 1); + x_164 = x_163; + lean_ctor_set_tag(x_164, 1); +} +lean_ctor_set(x_164, 0, x_139); +lean_ctor_set(x_164, 1, x_162); +return x_164; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_withNewRootScope(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_withNewRootScope___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_withNewRootScope___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: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_2); +lean_dec(x_2); +x_8 = l_Lean_Compiler_ToLCNF_withNewRootScope___rarg(x_1, x_7, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_expr_eqv(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +} +static size_t _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__1() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = 5; +x_3 = lean_usize_shift_left(x_1, x_2); +return x_3; +} +} +static size_t _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2() { +_start: +{ +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__1; +x_3 = lean_usize_sub(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t 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, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; lean_object* x_12; uint8_t 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_expr_eqv(x_3, x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_12); +x_14 = lean_box(0); +return x_14; +} +else +{ +lean_object* x_15; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_12); +return x_15; +} +} +case 1: +{ +lean_object* x_16; size_t x_17; +x_16 = lean_ctor_get(x_10, 0); +lean_inc(x_16); +lean_dec(x_10); +x_17 = lean_usize_shift_right(x_2, x_5); +x_1 = x_16; +x_2 = x_17; +goto _start; +} +default: +{ +lean_object* x_19; +x_19 = lean_box(0); +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_1, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 1); +lean_inc(x_21); +lean_dec(x_1); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__3(x_20, x_21, lean_box(0), x_22, x_3); +lean_dec(x_21); +lean_dec(x_20); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint64_t x_4; size_t x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_Expr_hash(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2(x_3, x_5, x_2); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6(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; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_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; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Expr_hash(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; } -lean_ctor_set(x_137, 0, x_118); -lean_ctor_set(x_137, 1, x_135); -return x_137; } } +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__7(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_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +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_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +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 +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_expr_eqv(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +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_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +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; +} +} +} +} +} +static lean_object* _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(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) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t 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_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___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); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_expr_eqv(x_4, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; +} +else +{ +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); +return x_1; +} +} +else +{ +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_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_expr_eqv(x_4, x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; +} +} +} +case 1: +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(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); +lean_ctor_set(x_1, 0, x_39); +return x_1; +} +else +{ +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(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); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; +} +} +default: +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} +} +} +} +else +{ +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___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); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); +} +x_63 = lean_expr_eqv(x_4, x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 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_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; +} +} +case 1: +{ +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); +} +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(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 { + x_76 = x_72; +} +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +default: +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +} +} +} +else +{ +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +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_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__7(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) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1; +x_93 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; +} +else +{ +return x_84; +} +} +else +{ +return x_84; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_1); +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_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__7(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) +{ +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1; +x_107 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; +} +else +{ +return x_98; +} +} +else +{ +return x_98; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_Expr_hash(x_2); +x_8 = lean_uint64_to_usize(x_7); +x_9 = 1; +x_10 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(x_5, x_8, x_9, x_2, x_3); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_10); +return x_1; +} +else +{ +lean_object* x_13; lean_object* x_14; uint64_t 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_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_1); +x_15 = l_Lean_Expr_hash(x_2); +x_16 = lean_uint64_to_usize(x_15); +x_17 = 1; +x_18 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(x_13, x_16, x_17, x_2, x_3); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_14, x_19); +lean_dec(x_14); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNFType(lean_object* x_1, uint8_t 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; uint8_t x_10; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_3, 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, 5); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_1); +x_14 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(x_13, x_1); +if (lean_obj_tag(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; 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_free_object(x_9); +x_15 = lean_st_ref_get(x_5, 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_3, 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); +x_21 = lean_box(0); +x_22 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; +x_23 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_20); +lean_ctor_set(x_25, 2, x_23); +lean_ctor_set(x_25, 3, x_21); +lean_ctor_set(x_25, 4, x_24); +lean_ctor_set(x_25, 5, x_21); +x_26 = lean_st_ref_get(x_5, x_19); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; +x_29 = lean_st_mk_ref(x_28, x_27); +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); +lean_inc(x_5); +lean_inc(x_30); +lean_inc(x_1); +x_32 = l_Lean_Compiler_toLCNFType(x_1, x_25, x_30, x_4, x_5, x_31); +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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_st_ref_get(x_5, x_34); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_get(x_30, x_36); +lean_dec(x_30); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_st_ref_get(x_5, x_38); +lean_dec(x_5); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_st_ref_take(x_3, x_40); +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_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_ctor_get(x_42, 5); +lean_inc(x_33); +x_46 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(x_45, x_1, x_33); +lean_ctor_set(x_42, 5, x_46); +x_47 = lean_st_ref_set(x_3, x_42, x_43); +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(x_47, 0, x_33); +return x_47; +} +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(0, 2, 0); +lean_ctor_set(x_51, 0, x_33); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +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; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_52 = lean_ctor_get(x_42, 0); +x_53 = lean_ctor_get(x_42, 1); +x_54 = lean_ctor_get(x_42, 2); +x_55 = lean_ctor_get(x_42, 3); +x_56 = lean_ctor_get(x_42, 4); +x_57 = lean_ctor_get(x_42, 5); +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_dec(x_42); +lean_inc(x_33); +x_58 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(x_57, x_1, x_33); +x_59 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_53); +lean_ctor_set(x_59, 2, x_54); +lean_ctor_set(x_59, 3, x_55); +lean_ctor_set(x_59, 4, x_56); +lean_ctor_set(x_59, 5, x_58); +x_60 = lean_st_ref_set(x_3, x_59, x_43); +x_61 = lean_ctor_get(x_60, 1); +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); +} +if (lean_is_scalar(x_62)) { + x_63 = lean_alloc_ctor(0, 2, 0); +} else { + x_63 = x_62; +} +lean_ctor_set(x_63, 0, x_33); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +else +{ +uint8_t x_64; +lean_dec(x_30); +lean_dec(x_5); +lean_dec(x_1); +x_64 = !lean_is_exclusive(x_32); +if (x_64 == 0) +{ +return x_32; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +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_32); +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_68; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_68 = lean_ctor_get(x_14, 0); +lean_inc(x_68); +lean_dec(x_14); +lean_ctor_set(x_9, 0, x_68); +return x_9; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_9, 0); +x_70 = lean_ctor_get(x_9, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_9); +x_71 = lean_ctor_get(x_69, 5); +lean_inc(x_71); +lean_dec(x_69); +lean_inc(x_1); +x_72 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNFType___spec__1(x_71, x_1); +if (lean_obj_tag(x_72) == 0) +{ +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_73 = lean_st_ref_get(x_5, x_70); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = lean_st_ref_get(x_3, 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, 0); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_box(0); +x_80 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; +x_81 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +x_82 = lean_unsigned_to_nat(0u); +x_83 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_78); +lean_ctor_set(x_83, 2, x_81); +lean_ctor_set(x_83, 3, x_79); +lean_ctor_set(x_83, 4, x_82); +lean_ctor_set(x_83, 5, x_79); +x_84 = lean_st_ref_get(x_5, x_77); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; +x_87 = lean_st_mk_ref(x_86, x_85); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +lean_inc(x_5); +lean_inc(x_88); +lean_inc(x_1); +x_90 = l_Lean_Compiler_toLCNFType(x_1, x_83, x_88, x_4, x_5, x_89); +if (lean_obj_tag(x_90) == 0) +{ +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; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_st_ref_get(x_5, x_92); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +lean_dec(x_93); +x_95 = lean_st_ref_get(x_88, x_94); +lean_dec(x_88); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = lean_st_ref_get(x_5, x_96); +lean_dec(x_5); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_99 = lean_st_ref_take(x_3, x_98); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_102 = lean_ctor_get(x_100, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_100, 1); +lean_inc(x_103); +x_104 = lean_ctor_get(x_100, 2); +lean_inc(x_104); +x_105 = lean_ctor_get(x_100, 3); +lean_inc(x_105); +x_106 = lean_ctor_get(x_100, 4); +lean_inc(x_106); +x_107 = lean_ctor_get(x_100, 5); +lean_inc(x_107); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + lean_ctor_release(x_100, 2); + lean_ctor_release(x_100, 3); + lean_ctor_release(x_100, 4); + lean_ctor_release(x_100, 5); + x_108 = x_100; +} else { + lean_dec_ref(x_100); + x_108 = lean_box(0); +} +lean_inc(x_91); +x_109 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNFType___spec__4(x_107, x_1, x_91); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 6, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_102); +lean_ctor_set(x_110, 1, x_103); +lean_ctor_set(x_110, 2, x_104); +lean_ctor_set(x_110, 3, x_105); +lean_ctor_set(x_110, 4, x_106); +lean_ctor_set(x_110, 5, x_109); +x_111 = lean_st_ref_set(x_3, x_110, x_101); +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_113 = x_111; +} else { + lean_dec_ref(x_111); + x_113 = lean_box(0); +} +if (lean_is_scalar(x_113)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_113; +} +lean_ctor_set(x_114, 0, x_91); +lean_ctor_set(x_114, 1, x_112); +return x_114; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_88); +lean_dec(x_5); +lean_dec(x_1); +x_115 = lean_ctor_get(x_90, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_90, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_117 = x_90; +} else { + lean_dec_ref(x_90); + x_117 = lean_box(0); +} +if (lean_is_scalar(x_117)) { + x_118 = lean_alloc_ctor(1, 2, 0); +} else { + x_118 = x_117; +} +lean_ctor_set(x_118, 0, x_115); +lean_ctor_set(x_118, 1, x_116); +return x_118; +} +} +else +{ +lean_object* x_119; lean_object* x_120; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_119 = lean_ctor_get(x_72, 0); +lean_inc(x_119); +lean_dec(x_72); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_70); +return x_120; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__3___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_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__3(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_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___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_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___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) { +_start: +{ +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNFType___spec__6(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_Compiler_ToLCNF_withNewRootScope(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_withNewRootScope___rarg___boxed), 6, 0); -return x_2; +size_t x_6; size_t x_7; lean_object* x_8; +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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_withNewRootScope___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_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNFType___boxed(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_2); lean_dec(x_2); -x_8 = l_Lean_Compiler_ToLCNF_withNewRootScope___rarg(x_1, x_7, x_3, x_4, x_5, x_6); +x_8 = l_Lean_Compiler_ToLCNF_toLCNFType(x_1, x_7, x_3, x_4, x_5, x_6); +lean_dec(x_3); return x_8; } } @@ -1616,189 +2960,146 @@ return x_10; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_mkLocalDecl(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) { _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; 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_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_9 = l_Lean_mkFreshFVarId___at_Lean_Compiler_ToLCNF_mkLocalDecl___spec__1(x_4, x_5, x_6, 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_st_ref_get(x_7, x_11); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_st_ref_get(x_5, x_13); -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(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_box(0); -x_19 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; -x_20 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_17); -lean_ctor_set(x_22, 2, x_20); -lean_ctor_set(x_22, 3, x_18); -lean_ctor_set(x_22, 4, x_21); -lean_ctor_set(x_22, 5, x_18); -x_23 = lean_st_ref_get(x_7, x_16); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; -x_26 = lean_st_mk_ref(x_25, x_24); -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_inc(x_7); -lean_inc(x_27); lean_inc(x_2); -x_29 = l_Lean_Compiler_toLCNFType(x_2, x_22, x_27, x_6, x_7, x_28); -if (lean_obj_tag(x_29) == 0) +x_12 = l_Lean_Compiler_ToLCNF_toLCNFType(x_2, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_12) == 0) { -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; uint8_t x_41; -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 = lean_st_ref_get(x_7, x_31); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = lean_st_ref_get(x_27, x_33); -lean_dec(x_27); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -x_36 = lean_st_ref_get(x_7, x_35); +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; +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_7, x_14); lean_dec(x_7); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_st_ref_take(x_5, x_37); -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_is_exclusive(x_39); -if (x_41 == 0) +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_take(x_5, 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_is_exclusive(x_18); +if (x_20 == 0) { -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_42 = lean_ctor_get(x_39, 0); -x_43 = lean_ctor_get(x_39, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_18, 0); +x_22 = lean_ctor_get(x_18, 1); lean_inc(x_1); lean_inc(x_10); -x_44 = lean_local_ctx_mk_local_decl(x_42, x_10, x_1, x_2, x_3); +x_23 = lean_local_ctx_mk_local_decl(x_21, x_10, x_1, x_2, x_3); lean_inc(x_10); -x_45 = lean_local_ctx_mk_local_decl(x_43, x_10, x_1, x_30, x_3); -lean_ctor_set(x_39, 1, x_45); -lean_ctor_set(x_39, 0, x_44); -x_46 = lean_st_ref_set(x_5, x_39, x_40); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_24 = lean_local_ctx_mk_local_decl(x_22, x_10, x_1, x_13, x_3); +lean_ctor_set(x_18, 1, x_24); +lean_ctor_set(x_18, 0, x_23); +x_25 = lean_st_ref_set(x_5, x_18, x_19); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_46, 0); -lean_dec(x_48); -x_49 = l_Lean_Expr_fvar___override(x_10); -lean_ctor_set(x_46, 0, x_49); -return x_46; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +x_28 = l_Lean_Expr_fvar___override(x_10); +lean_ctor_set(x_25, 0, x_28); +return x_25; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_46, 1); -lean_inc(x_50); -lean_dec(x_46); -x_51 = l_Lean_Expr_fvar___override(x_10); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_dec(x_25); +x_30 = l_Lean_Expr_fvar___override(x_10); +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_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_39, 0); -x_54 = lean_ctor_get(x_39, 1); -x_55 = lean_ctor_get(x_39, 2); -x_56 = lean_ctor_get(x_39, 3); -x_57 = lean_ctor_get(x_39, 4); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_39); +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_32 = lean_ctor_get(x_18, 0); +x_33 = lean_ctor_get(x_18, 1); +x_34 = lean_ctor_get(x_18, 2); +x_35 = lean_ctor_get(x_18, 3); +x_36 = lean_ctor_get(x_18, 4); +x_37 = lean_ctor_get(x_18, 5); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_18); lean_inc(x_1); lean_inc(x_10); -x_58 = lean_local_ctx_mk_local_decl(x_53, x_10, x_1, x_2, x_3); +x_38 = lean_local_ctx_mk_local_decl(x_32, x_10, x_1, x_2, x_3); lean_inc(x_10); -x_59 = lean_local_ctx_mk_local_decl(x_54, x_10, x_1, x_30, x_3); -x_60 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_55); -lean_ctor_set(x_60, 3, x_56); -lean_ctor_set(x_60, 4, x_57); -x_61 = lean_st_ref_set(x_5, x_60, x_40); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_63 = x_61; +x_39 = lean_local_ctx_mk_local_decl(x_33, x_10, x_1, x_13, x_3); +x_40 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_34); +lean_ctor_set(x_40, 3, x_35); +lean_ctor_set(x_40, 4, x_36); +lean_ctor_set(x_40, 5, x_37); +x_41 = lean_st_ref_set(x_5, x_40, x_19); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_43 = x_41; } else { - lean_dec_ref(x_61); - x_63 = lean_box(0); + lean_dec_ref(x_41); + x_43 = lean_box(0); } -x_64 = l_Lean_Expr_fvar___override(x_10); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); +x_44 = l_Lean_Expr_fvar___override(x_10); +if (lean_is_scalar(x_43)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - x_65 = x_63; + x_45 = x_43; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; } } else { -uint8_t x_66; -lean_dec(x_27); +uint8_t x_46; lean_dec(x_10); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_66 = !lean_is_exclusive(x_29); -if (x_66 == 0) +x_46 = !lean_is_exclusive(x_12); +if (x_46 == 0) { -return x_29; +return x_12; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_29, 0); -x_68 = lean_ctor_get(x_29, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_29); -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_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_12, 0); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_12); +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; } } } @@ -1911,46 +3212,49 @@ 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; lean_object* x_37; lean_object* x_38; lean_object* x_39; +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_28 = lean_ctor_get(x_17, 0); x_29 = lean_ctor_get(x_17, 1); x_30 = lean_ctor_get(x_17, 2); x_31 = lean_ctor_get(x_17, 3); x_32 = lean_ctor_get(x_17, 4); +x_33 = lean_ctor_get(x_17, 5); +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_dec(x_17); -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_31, x_33); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_31, x_34); lean_dec(x_31); -x_35 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_35, 0, x_28); -lean_ctor_set(x_35, 1, x_29); -lean_ctor_set(x_35, 2, x_30); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_32); -x_36 = lean_st_ref_set(x_3, x_35, x_18); -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; +x_36 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_36, 0, x_28); +lean_ctor_set(x_36, 1, x_29); +lean_ctor_set(x_36, 2, x_30); +lean_ctor_set(x_36, 3, x_35); +lean_ctor_set(x_36, 4, x_32); +lean_ctor_set(x_36, 5, x_33); +x_37 = lean_st_ref_set(x_3, x_36, x_18); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; } else { - lean_dec_ref(x_36); - x_38 = lean_box(0); + lean_dec_ref(x_37); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_13); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_13); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } } @@ -2025,51 +3329,54 @@ return x_31; } 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; lean_object* x_38; uint8_t 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_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; uint8_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; x_32 = lean_ctor_get(x_16, 0); x_33 = lean_ctor_get(x_16, 1); x_34 = lean_ctor_get(x_16, 2); x_35 = lean_ctor_get(x_16, 3); x_36 = lean_ctor_get(x_16, 4); +x_37 = lean_ctor_get(x_16, 5); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_16); -x_37 = 0; +x_38 = 0; lean_inc(x_6); lean_inc(x_1); -x_38 = lean_local_ctx_mk_let_decl(x_32, x_1, x_6, x_2, x_3, x_37); -x_39 = 1; -x_40 = lean_local_ctx_mk_let_decl(x_33, x_1, x_6, x_4, x_5, x_39); +x_39 = lean_local_ctx_mk_let_decl(x_32, x_1, x_6, x_2, x_3, x_38); +x_40 = 1; +x_41 = lean_local_ctx_mk_let_decl(x_33, x_1, x_6, x_4, x_5, x_40); lean_inc(x_12); -x_41 = lean_array_push(x_34, x_12); -x_42 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_42, 0, x_38); -lean_ctor_set(x_42, 1, x_40); -lean_ctor_set(x_42, 2, x_41); -lean_ctor_set(x_42, 3, x_35); -lean_ctor_set(x_42, 4, x_36); -x_43 = lean_st_ref_set(x_8, x_42, x_17); -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -if (lean_is_exclusive(x_43)) { - lean_ctor_release(x_43, 0); - lean_ctor_release(x_43, 1); - x_45 = x_43; +x_42 = lean_array_push(x_34, x_12); +x_43 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_41); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set(x_43, 3, x_35); +lean_ctor_set(x_43, 4, x_36); +lean_ctor_set(x_43, 5, x_37); +x_44 = lean_st_ref_set(x_8, x_43, x_17); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; } else { - lean_dec_ref(x_43); - x_45 = lean_box(0); + lean_dec_ref(x_44); + x_46 = lean_box(0); } -if (lean_is_scalar(x_45)) { - x_46 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); } else { - x_46 = x_45; + x_47 = x_46; } -lean_ctor_set(x_46, 0, x_12); -lean_ctor_set(x_46, 1, x_44); -return x_46; +lean_ctor_set(x_47, 0, x_12); +lean_ctor_set(x_47, 1, x_45); +return x_47; } } } @@ -2247,85 +3554,88 @@ return x_39; } else { -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_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; x_40 = lean_ctor_get(x_27, 0); x_41 = lean_ctor_get(x_27, 1); x_42 = lean_ctor_get(x_27, 2); x_43 = lean_ctor_get(x_27, 3); x_44 = lean_ctor_get(x_27, 4); +x_45 = lean_ctor_get(x_27, 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_27); -x_45 = 1; -x_46 = lean_local_ctx_mk_let_decl(x_41, x_8, x_12, x_21, x_1, x_45); +x_46 = 1; +x_47 = lean_local_ctx_mk_let_decl(x_41, x_8, x_12, x_21, x_1, x_46); lean_inc(x_23); -x_47 = lean_array_push(x_42, x_23); -x_48 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_48, 0, x_40); -lean_ctor_set(x_48, 1, x_46); -lean_ctor_set(x_48, 2, x_47); -lean_ctor_set(x_48, 3, x_43); -lean_ctor_set(x_48, 4, x_44); -x_49 = lean_st_ref_set(x_3, x_48, x_28); -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; +x_48 = lean_array_push(x_42, x_23); +x_49 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_49, 0, x_40); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_48); +lean_ctor_set(x_49, 3, x_43); +lean_ctor_set(x_49, 4, x_44); +lean_ctor_set(x_49, 5, x_45); +x_50 = lean_st_ref_set(x_3, x_49, x_28); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; } else { - lean_dec_ref(x_49); - x_51 = lean_box(0); + lean_dec_ref(x_50); + x_52 = lean_box(0); } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_51; + x_53 = x_52; } -lean_ctor_set(x_52, 0, x_23); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_ctor_set(x_53, 0, x_23); +lean_ctor_set(x_53, 1, x_51); +return x_53; } } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_12); lean_dec(x_8); lean_dec(x_5); lean_dec(x_1); -x_53 = !lean_is_exclusive(x_20); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_20); +if (x_54 == 0) { return x_20; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_20, 0); -x_55 = lean_ctor_get(x_20, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_20, 0); +x_56 = lean_ctor_get(x_20, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); lean_dec(x_20); -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; +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 { -lean_object* x_57; +lean_object* x_58; lean_dec(x_5); lean_dec(x_4); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_1); -lean_ctor_set(x_57, 1, x_6); -return x_57; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_1); +lean_ctor_set(x_58, 1, x_6); +return x_58; } } } @@ -3165,27 +4475,53 @@ return x_6; } else { -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_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; lean_object* x_19; lean_object* x_20; +lean_object* x_9; lean_object* x_10; size_t x_11; size_t 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; uint64_t x_19; uint8_t x_20; x_9 = lean_array_fget(x_2, x_5); x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Expr_hash(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -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); +x_11 = 1; +x_12 = lean_usize_sub(x_1, x_11); +x_13 = 5; +x_14 = lean_usize_mul(x_13, x_12); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_5, x_15); lean_dec(x_5); -x_20 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_6, x_17, x_1, x_9, x_10); +x_17 = lean_ctor_get(x_9, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +x_19 = l_Lean_Expr_hash(x_17); +lean_dec(x_17); +x_20 = lean_unbox(x_18); +lean_dec(x_18); +if (x_20 == 0) +{ +uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; lean_object* x_25; +x_21 = 13; +x_22 = lean_uint64_mix_hash(x_19, x_21); +x_23 = lean_uint64_to_usize(x_22); +x_24 = lean_usize_shift_right(x_23, x_14); +x_25 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_6, x_24, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; +x_5 = x_16; +x_6 = x_25; +goto _start; +} +else +{ +uint64_t x_27; uint64_t x_28; size_t x_29; size_t x_30; lean_object* x_31; +x_27 = 11; +x_28 = lean_uint64_mix_hash(x_19, x_27); +x_29 = lean_uint64_to_usize(x_28); +x_30 = lean_usize_shift_right(x_29, x_14); +x_31 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_6, x_30, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_31; goto _start; } } } +} LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -3229,81 +4565,102 @@ return x_16; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_expr_eqv(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_17; 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_array_fget(x_5, x_2); +x_23 = lean_ctor_get(x_3, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_3, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_expr_eqv(x_23, x_25); +lean_dec(x_25); +lean_dec(x_23); +if (x_27 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); +lean_dec(x_24); lean_dec(x_6); lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_add(x_2, x_28); lean_dec(x_2); -x_2 = x_20; +x_2 = x_29; goto _start; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -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_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); +uint8_t x_31; +x_31 = lean_unbox(x_24); lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +if (x_31 == 0) +{ +uint8_t x_32; +x_32 = lean_unbox(x_26); +lean_dec(x_26); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_1); +x_33 = lean_box(0); +x_17 = x_33; +goto block_21; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_34; lean_object* x_35; +lean_dec(x_6); +lean_dec(x_5); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_2, x_34); lean_dec(x_2); -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; -} -} +x_2 = x_35; +goto _start; } } +else +{ +uint8_t x_37; +x_37 = lean_unbox(x_26); +lean_dec(x_26); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_6); +lean_dec(x_5); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_add(x_2, x_38); +lean_dec(x_2); +x_2 = x_39; +goto _start; } -static size_t _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__1() { -_start: +else { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = 5; -x_3 = lean_usize_shift_left(x_1, x_2); -return x_3; +lean_object* x_41; +lean_dec(x_1); +x_41 = lean_box(0); +x_17 = x_41; +goto block_21; } } -static size_t _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2() { -_start: +} +block_21: { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__1; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_17); +x_18 = lean_array_fset(x_5, x_2, x_3); +x_19 = lean_array_fset(x_6, x_2, x_4); +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; } } -static lean_object* _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = l_Std_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); -return x_1; } } LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { @@ -3319,7 +4676,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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2; +x_10 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___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); @@ -3345,113 +4702,244 @@ uint8_t x_18; x_18 = !lean_is_exclusive(x_15); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_19 = lean_ctor_get(x_15, 0); x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_expr_eqv(x_4, x_19); -if (x_21 == 0) +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +x_25 = lean_expr_eqv(x_21, x_23); +lean_dec(x_23); +lean_dec(x_21); +if (x_25 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_24); +lean_dec(x_22); lean_free_object(x_15); -x_22 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); +x_26 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_array_fset(x_17, x_12, x_27); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); +lean_ctor_set(x_1, 0, x_28); return x_1; } else { -lean_object* x_25; +uint8_t x_29; +x_29 = lean_unbox(x_22); +lean_dec(x_22); +if (x_29 == 0) +{ +uint8_t x_30; +x_30 = lean_unbox(x_24); +lean_dec(x_24); +if (x_30 == 0) +{ +lean_object* x_31; lean_dec(x_20); lean_dec(x_19); lean_ctor_set(x_15, 1, x_5); lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); +x_31 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); +lean_ctor_set(x_1, 0, x_31); +return x_1; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_free_object(x_15); +x_32 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_array_fset(x_17, x_12, x_33); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_34); return x_1; } } else { -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_15, 1); -lean_inc(x_27); -lean_inc(x_26); +uint8_t x_35; +x_35 = lean_unbox(x_24); +lean_dec(x_24); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_free_object(x_15); +x_36 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_array_fset(x_17, x_12, x_37); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_38); +return x_1; +} +else +{ +lean_object* x_39; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; +} +} +} +} +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; uint8_t x_46; +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_28 = lean_expr_eqv(x_4, x_26); -if (x_28 == 0) +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_4, 1); +lean_inc(x_43); +x_44 = lean_ctor_get(x_40, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +x_46 = lean_expr_eqv(x_42, x_44); +lean_dec(x_44); +lean_dec(x_42); +if (x_46 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_40, x_41, x_4, x_5); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_array_fset(x_17, x_12, x_48); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); +lean_ctor_set(x_1, 0, x_49); return x_1; } else { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); +uint8_t x_50; +x_50 = lean_unbox(x_43); +lean_dec(x_43); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = lean_unbox(x_45); +lean_dec(x_45); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_41); +lean_dec(x_40); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_4); +lean_ctor_set(x_52, 1, x_5); +x_53 = lean_array_fset(x_17, x_12, x_52); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); +lean_ctor_set(x_1, 0, x_53); +return x_1; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_40, x_41, x_4, x_5); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_array_fset(x_17, x_12, x_55); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_56); +return x_1; +} +} +else +{ +uint8_t x_57; +x_57 = lean_unbox(x_45); +lean_dec(x_45); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_40, x_41, x_4, x_5); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_60 = lean_array_fset(x_17, x_12, x_59); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_60); +return x_1; +} +else +{ +lean_object* x_61; lean_object* x_62; +lean_dec(x_41); +lean_dec(x_40); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_4); +lean_ctor_set(x_61, 1, x_5); +x_62 = lean_array_fset(x_17, x_12, x_61); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_62); return x_1; } } } +} +} case 1: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +uint8_t x_63; +x_63 = !lean_is_exclusive(x_15); +if (x_63 == 0) { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(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_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; +x_64 = lean_ctor_get(x_15, 0); +x_65 = lean_usize_shift_right(x_2, x_9); +x_66 = lean_usize_add(x_3, x_8); +x_67 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_64, x_65, x_66, x_4, x_5); +lean_ctor_set(x_15, 0, x_67); +x_68 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); +lean_ctor_set(x_1, 0, x_68); return x_1; } else { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); +lean_object* x_69; size_t x_70; size_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_69 = lean_ctor_get(x_15, 0); +lean_inc(x_69); 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_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(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); +x_70 = lean_usize_shift_right(x_2, x_9); +x_71 = lean_usize_add(x_3, x_8); +x_72 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_69, x_70, x_71, x_4, x_5); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_72); +x_74 = lean_array_fset(x_17, x_12, x_73); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); +lean_ctor_set(x_1, 0, x_74); return x_1; } } default: { -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); +lean_object* x_75; lean_object* x_76; +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_4); +lean_ctor_set(x_75, 1, x_5); +x_76 = lean_array_fset(x_17, x_12, x_75); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); +lean_ctor_set(x_1, 0, x_76); return x_1; } } @@ -3459,121 +4947,195 @@ return x_1; } else { -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); +lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; size_t x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_77 = lean_ctor_get(x_1, 0); +lean_inc(x_77); lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___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); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +x_78 = 1; +x_79 = 5; +x_80 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2; +x_81 = lean_usize_land(x_2, x_80); +x_82 = lean_usize_to_nat(x_81); +x_83 = lean_array_get_size(x_77); +x_84 = lean_nat_dec_lt(x_82, x_83); +lean_dec(x_83); +if (x_84 == 0) { -lean_object* x_56; -lean_dec(x_53); +lean_object* x_85; +lean_dec(x_82); lean_dec(x_5); lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; +x_85 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_85, 0, x_77); +return x_85; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_array_fget(x_77, x_82); +x_87 = lean_box(0); +x_88 = lean_array_fset(x_77, x_82, x_87); +switch (lean_obj_tag(x_86)) { +case 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; uint8_t x_96; +x_89 = lean_ctor_get(x_86, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_86, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_91 = x_86; +} else { + lean_dec_ref(x_86); + x_91 = lean_box(0); +} +x_92 = lean_ctor_get(x_4, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_4, 1); +lean_inc(x_93); +x_94 = lean_ctor_get(x_89, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_89, 1); +lean_inc(x_95); +x_96 = lean_expr_eqv(x_92, x_94); +lean_dec(x_94); +lean_dec(x_92); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_95); +lean_dec(x_93); +lean_dec(x_91); +x_97 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_89, x_90, x_4, x_5); +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_97); +x_99 = lean_array_fset(x_88, x_82, x_98); +lean_dec(x_82); +x_100 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_100, 0, x_99); +return x_100; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: +uint8_t x_101; +x_101 = lean_unbox(x_93); +lean_dec(x_93); +if (x_101 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; +uint8_t x_102; +x_102 = lean_unbox(x_95); +lean_dec(x_95); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_90); +lean_dec(x_89); +if (lean_is_scalar(x_91)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_57); - x_62 = lean_box(0); + x_103 = x_91; } -x_63 = lean_expr_eqv(x_4, x_60); -if (x_63 == 0) +lean_ctor_set(x_103, 0, x_4); +lean_ctor_set(x_103, 1, x_5); +x_104 = lean_array_fset(x_88, x_82, x_103); +lean_dec(x_82); +x_105 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_105, 0, x_104); +return x_105; +} +else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_91); +x_106 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_89, x_90, x_4, x_5); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = lean_array_fset(x_88, x_82, x_107); +lean_dec(x_82); +x_109 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_109, 0, x_108); +return x_109; +} } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); +uint8_t x_110; +x_110 = lean_unbox(x_95); +lean_dec(x_95); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_91); +x_111 = l_Std_PersistentHashMap_mkCollisionNode___rarg(x_89, x_90, x_4, x_5); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = lean_array_fset(x_88, x_82, x_112); +lean_dec(x_82); +x_114 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_114, 0, x_113); +return x_114; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_90); +lean_dec(x_89); +if (lean_is_scalar(x_91)) { + x_115 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_62; + x_115 = x_91; +} +lean_ctor_set(x_115, 0, x_4); +lean_ctor_set(x_115, 1, x_5); +x_116 = lean_array_fset(x_88, x_82, x_115); +lean_dec(x_82); +x_117 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_117, 0, x_116); +return x_117; +} } -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; } } case 1: { -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; +lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_118 = lean_ctor_get(x_86, 0); +lean_inc(x_118); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + x_119 = x_86; } else { - lean_dec_ref(x_57); - x_72 = lean_box(0); -} -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); + lean_dec_ref(x_86); + x_119 = lean_box(0); +} +x_120 = lean_usize_shift_right(x_2, x_79); +x_121 = lean_usize_add(x_3, x_78); +x_122 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_118, x_120, x_121, x_4, x_5); +if (lean_is_scalar(x_119)) { + x_123 = lean_alloc_ctor(1, 1, 0); } else { - x_76 = x_72; + x_123 = x_119; } -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; +lean_ctor_set(x_123, 0, x_122); +x_124 = lean_array_fset(x_88, x_82, x_123); +lean_dec(x_82); +x_125 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_125, 0, x_124); +return x_125; } default: { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_4); +lean_ctor_set(x_126, 1, x_5); +x_127 = lean_array_fset(x_88, x_82, x_126); +lean_dec(x_82); +x_128 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_128, 0, x_127); +return x_128; } } } @@ -3581,90 +5143,90 @@ return x_81; } else { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) +uint8_t x_129; +x_129 = !lean_is_exclusive(x_1); +if (x_129 == 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_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(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) +lean_object* x_130; lean_object* x_131; size_t x_132; uint8_t x_133; +x_130 = lean_unsigned_to_nat(0u); +x_131 = l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(x_1, x_130, x_4, x_5); +x_132 = 7; +x_133 = lean_usize_dec_le(x_132, x_3); +if (x_133 == 0) { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_134 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_131); +x_135 = lean_unsigned_to_nat(4u); +x_136 = lean_nat_dec_lt(x_134, x_135); +lean_dec(x_134); +if (x_136 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3; -x_93 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_137 = lean_ctor_get(x_131, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_131, 1); +lean_inc(x_138); +lean_dec(x_131); +x_139 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1; +x_140 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__3(x_3, x_137, x_138, lean_box(0), x_130, x_139); +lean_dec(x_138); +lean_dec(x_137); +return x_140; } else { -return x_84; +return x_131; } } else { -return x_84; +return x_131; } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; size_t x_146; uint8_t x_147; +x_141 = lean_ctor_get(x_1, 0); +x_142 = lean_ctor_get(x_1, 1); +lean_inc(x_142); +lean_inc(x_141); lean_dec(x_1); -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_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(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) -{ -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_unsigned_to_nat(0u); +x_145 = l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__4(x_143, x_144, x_4, x_5); +x_146 = 7; +x_147 = lean_usize_dec_le(x_146, x_3); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; uint8_t x_150; +x_148 = l_Std_PersistentHashMap_getCollisionNodeSize___rarg(x_145); +x_149 = lean_unsigned_to_nat(4u); +x_150 = lean_nat_dec_lt(x_148, x_149); +lean_dec(x_148); +if (x_150 == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3; -x_107 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_151 = lean_ctor_get(x_145, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_145, 1); +lean_inc(x_152); +lean_dec(x_145); +x_153 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1; +x_154 = l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__3(x_3, x_151, x_152, lean_box(0), x_144, x_153); +lean_dec(x_152); +lean_dec(x_151); +return x_154; } else { -return x_98; +return x_145; } } else { -return x_98; +return x_145; } } } @@ -3677,39 +5239,88 @@ uint8_t x_4; x_4 = !lean_is_exclusive(x_1); if (x_4 == 0) { -lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; uint8_t x_13; x_5 = lean_ctor_get(x_1, 0); x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_Expr_hash(x_2); -x_8 = lean_uint64_to_usize(x_7); -x_9 = 1; -x_10 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_5, x_8, x_9, x_2, x_3); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_6, x_11); +x_7 = 1; +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_add(x_6, x_8); lean_dec(x_6); -lean_ctor_set(x_1, 1, x_12); -lean_ctor_set(x_1, 0, x_10); +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +x_12 = l_Lean_Expr_hash(x_10); +lean_dec(x_10); +x_13 = lean_unbox(x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +uint64_t x_14; uint64_t x_15; size_t x_16; lean_object* x_17; +x_14 = 13; +x_15 = lean_uint64_mix_hash(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_5, x_16, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_17); return x_1; } else { -lean_object* x_13; lean_object* x_14; uint64_t 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_13 = lean_ctor_get(x_1, 0); -x_14 = lean_ctor_get(x_1, 1); -lean_inc(x_14); -lean_inc(x_13); +uint64_t x_18; uint64_t x_19; size_t x_20; lean_object* x_21; +x_18 = 11; +x_19 = lean_uint64_mix_hash(x_12, x_18); +x_20 = lean_uint64_to_usize(x_19); +x_21 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_5, x_20, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_21); +return x_1; +} +} +else +{ +lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint8_t x_30; +x_22 = lean_ctor_get(x_1, 0); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +lean_inc(x_22); lean_dec(x_1); -x_15 = l_Lean_Expr_hash(x_2); -x_16 = lean_uint64_to_usize(x_15); -x_17 = 1; -x_18 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_13, x_16, x_17, x_2, x_3); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_14, x_19); -lean_dec(x_14); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_20); -return x_21; +x_24 = 1; +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_add(x_23, x_25); +lean_dec(x_23); +x_27 = lean_ctor_get(x_2, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 1); +lean_inc(x_28); +x_29 = l_Lean_Expr_hash(x_27); +lean_dec(x_27); +x_30 = lean_unbox(x_28); +lean_dec(x_28); +if (x_30 == 0) +{ +uint64_t x_31; uint64_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; +x_31 = 13; +x_32 = lean_uint64_mix_hash(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_22, x_33, x_24, x_2, x_3); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_26); +return x_35; +} +else +{ +uint64_t x_36; uint64_t x_37; size_t x_38; lean_object* x_39; lean_object* x_40; +x_36 = 11; +x_37 = lean_uint64_mix_hash(x_29, x_36); +x_38 = lean_uint64_to_usize(x_37); +x_39 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2(x_22, x_38, x_24, x_2, x_3); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_26); +return x_40; +} } } } @@ -3791,44 +5402,98 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__9(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_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; +} +else +{ +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_array_fget(x_1, x_4); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_5, 1); +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_expr_eqv(x_10, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_13); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_4, x_15); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_16; +goto _start; +} +else { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) +uint8_t x_18; +x_18 = lean_unbox(x_11); +if (x_18 == 0) { -lean_object* x_8; +uint8_t x_19; +x_19 = lean_unbox(x_13); +lean_dec(x_13); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_array_fget(x_2, x_4); lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_expr_eqv(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_4, x_22); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_23; +goto _start; +} +} +else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); +uint8_t x_25; +x_25 = lean_unbox(x_13); +lean_dec(x_13); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_4, x_26); lean_dec(x_4); x_3 = lean_box(0); -x_4 = x_12; +x_4 = x_27; goto _start; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); +lean_object* x_29; lean_object* x_30; +x_29 = lean_array_fget(x_2, x_4); lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +return x_30; +} +} } } } @@ -3843,7 +5508,7 @@ x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); lean_dec(x_1); x_5 = 5; -x_6 = l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2; +x_6 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2; x_7 = lean_usize_land(x_2, x_6); x_8 = lean_usize_to_nat(x_7); x_9 = lean_box(2); @@ -3853,76 +5518,145 @@ lean_dec(x_4); switch (lean_obj_tag(x_10)) { case 0: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; +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; 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_expr_eqv(x_3, x_11); +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_11, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); lean_dec(x_11); -if (x_13 == 0) +x_17 = lean_expr_eqv(x_13, x_15); +lean_dec(x_15); +if (x_17 == 0) { -lean_object* x_14; +lean_object* x_18; +lean_dec(x_16); lean_dec(x_12); -x_14 = lean_box(0); -return x_14; +x_18 = lean_box(0); +return x_18; } else { -lean_object* x_15; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_12); -return x_15; +uint8_t x_19; +x_19 = lean_unbox(x_14); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = lean_unbox(x_16); +lean_dec(x_16); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_12); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_12); +x_22 = lean_box(0); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = lean_unbox(x_16); +lean_dec(x_16); +if (x_23 == 0) +{ +lean_object* x_24; +lean_dec(x_12); +x_24 = lean_box(0); +return x_24; +} +else +{ +lean_object* x_25; +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_12); +return x_25; +} +} } } case 1: { -lean_object* x_16; size_t x_17; -x_16 = lean_ctor_get(x_10, 0); -lean_inc(x_16); +lean_object* x_26; size_t x_27; +x_26 = lean_ctor_get(x_10, 0); +lean_inc(x_26); lean_dec(x_10); -x_17 = lean_usize_shift_right(x_2, x_5); -x_1 = x_16; -x_2 = x_17; +x_27 = lean_usize_shift_right(x_2, x_5); +x_1 = x_26; +x_2 = x_27; goto _start; } default: { -lean_object* x_19; -x_19 = lean_box(0); -return x_19; +lean_object* x_29; +x_29 = lean_box(0); +return x_29; } } } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 1); -lean_inc(x_21); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_1, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 1); +lean_inc(x_31); lean_dec(x_1); -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__9(x_20, x_21, lean_box(0), x_22, x_3); -lean_dec(x_21); -lean_dec(x_20); -return x_23; +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__9(x_30, x_31, lean_box(0), x_32, x_3); +lean_dec(x_31); +lean_dec(x_30); +return x_33; } } } LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; uint64_t x_4; size_t x_5; lean_object* x_6; +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; uint8_t x_7; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); -x_4 = l_Lean_Expr_hash(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__8(x_3, x_5, x_2); +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 = l_Lean_Expr_hash(x_4); +lean_dec(x_4); +x_7 = lean_unbox(x_5); +lean_dec(x_5); +if (x_7 == 0) +{ +uint64_t x_8; uint64_t x_9; size_t x_10; lean_object* x_11; +x_8 = 13; +x_9 = lean_uint64_mix_hash(x_6, x_8); +x_10 = lean_uint64_to_usize(x_9); +x_11 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__8(x_3, x_10, x_2); lean_dec(x_2); -return x_6; +return x_11; +} +else +{ +uint64_t x_12; uint64_t x_13; size_t x_14; lean_object* x_15; +x_12 = 11; +x_13 = lean_uint64_mix_hash(x_6, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__8(x_3, x_14, x_2); +lean_dec(x_2); +return x_15; +} } } static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__10___closed__1() { @@ -4005,12 +5739,14 @@ return x_20; } 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_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_21 = lean_ctor_get(x_11, 0); x_22 = lean_ctor_get(x_11, 1); x_23 = lean_ctor_get(x_11, 2); x_24 = lean_ctor_get(x_11, 3); x_25 = lean_ctor_get(x_11, 4); +x_26 = lean_ctor_get(x_11, 5); +lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); @@ -4018,32 +5754,33 @@ lean_inc(x_22); lean_inc(x_21); lean_dec(x_11); lean_inc(x_2); -x_26 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__1(x_25, x_1, x_2); -x_27 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_27, 0, x_21); -lean_ctor_set(x_27, 1, x_22); -lean_ctor_set(x_27, 2, x_23); -lean_ctor_set(x_27, 3, x_24); -lean_ctor_set(x_27, 4, x_26); -x_28 = lean_st_ref_set(x_4, x_27, x_12); -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 = l_Std_PersistentHashMap_insert___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__1(x_25, x_1, x_2); +x_28 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_22); +lean_ctor_set(x_28, 2, x_23); +lean_ctor_set(x_28, 3, x_24); +lean_ctor_set(x_28, 4, x_27); +lean_ctor_set(x_28, 5, x_26); +x_29 = lean_st_ref_set(x_4, x_28, x_12); +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); } -if (lean_is_scalar(x_30)) { - x_31 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_31)) { + x_32 = lean_alloc_ctor(0, 2, 0); } else { - x_31 = x_30; + x_32 = x_31; } -lean_ctor_set(x_31, 0, x_2); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_ctor_set(x_32, 0, x_2); +lean_ctor_set(x_32, 1, x_30); +return x_32; } } } @@ -4077,7 +5814,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitCore___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; x_1 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(185u); +x_3 = lean_unsigned_to_nat(196u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4124,509 +5861,508 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitCore___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; x_1 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(183u); +x_3 = lean_unsigned_to_nat(194u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___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_Compiler_ToLCNF_toLCNF_visitCore___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_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(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) { _start: { -lean_dec(x_2); -switch (lean_obj_tag(x_1)) { +lean_dec(x_3); +switch (lean_obj_tag(x_2)) { case 0: { -lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__4; +lean_object* x_9; lean_object* x_10; +lean_dec(x_2); +x_9 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__4; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_9 = l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5(x_8, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_9) == 0) +x_10 = l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5(x_9, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_11, 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); -return x_12; +return x_13; } else { -uint8_t x_13; +uint8_t x_14; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_13 = !lean_is_exclusive(x_9); -if (x_13 == 0) +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) { -return x_9; +return x_10; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_9, 0); -x_15 = lean_ctor_get(x_9, 1); +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_inc(x_14); -lean_dec(x_9); -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_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; } } } case 2: { -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_17 = l_Lean_indentExpr(x_1); -x_18 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__6; -x_19 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__8; -x_21 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_throwError___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__6(x_21, x_3, x_4, x_5, x_6, x_7); +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; +lean_dec(x_1); +x_18 = l_Lean_indentExpr(x_2); +x_19 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__6; +x_20 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +x_21 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__8; +x_22 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_throwError___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__6(x_22, 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); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -return x_22; +return x_23; } else { -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_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_inc(x_24); -lean_dec(x_22); -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_dec(x_23); +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; } } case 4: { -lean_object* x_27; +lean_object* x_28; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_27 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp(x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_28) == 0) { -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_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_28, x_3, x_4, x_5, x_6, x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_29, x_4, x_5, x_6, x_7, x_30); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_30; +lean_dec(x_5); +return x_31; } else { -uint8_t x_31; +uint8_t x_32; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_27); -if (x_31 == 0) +x_32 = !lean_is_exclusive(x_28); +if (x_32 == 0) { -return x_27; +return x_28; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_27, 0); -x_33 = lean_ctor_get(x_27, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_28, 0); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_27); -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_dec(x_28); +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; } } } case 5: { -lean_object* x_35; +lean_object* x_36; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_35 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_35) == 0) +x_36 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp(x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_36) == 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); -x_37 = lean_ctor_get(x_35, 1); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_36, x_3, x_4, x_5, x_6, x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_37, x_4, x_5, x_6, x_7, x_38); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_38; +return x_39; } else { -uint8_t x_39; +uint8_t x_40; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_35); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_36); +if (x_40 == 0) { -return x_35; +return x_36; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_35, 0); -x_41 = lean_ctor_get(x_35, 1); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_36, 0); +x_42 = lean_ctor_get(x_36, 1); +lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_35); -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_dec(x_36); +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; } } } case 6: { -lean_object* x_43; +lean_object* x_44; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_43 = l_Lean_Compiler_ToLCNF_toLCNF_visitLambda(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_43) == 0) +x_44 = l_Lean_Compiler_ToLCNF_toLCNF_visitLambda(x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_44, x_3, x_4, x_5, x_6, x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_45, x_4, x_5, x_6, x_7, x_46); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_46; +return x_47; } else { -uint8_t x_47; +uint8_t x_48; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_47 = !lean_is_exclusive(x_43); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_44); +if (x_48 == 0) { -return x_43; +return x_44; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_43, 0); -x_49 = lean_ctor_get(x_43, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_44, 0); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_43); -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_44); +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; } } } case 7: { -lean_object* x_51; lean_object* x_52; -x_51 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__9; +lean_object* x_52; lean_object* x_53; +lean_dec(x_2); +x_52 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__9; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_52 = l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5(x_51, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_52) == 0) +x_53 = l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5(x_52, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_53) == 0) { -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_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); -lean_dec(x_52); -x_55 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_53, x_3, x_4, x_5, x_6, x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_54, x_4, x_5, x_6, x_7, x_55); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_55; +return x_56; } else { -uint8_t x_56; +uint8_t x_57; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_56 = !lean_is_exclusive(x_52); -if (x_56 == 0) +x_57 = !lean_is_exclusive(x_53); +if (x_57 == 0) { -return x_52; +return x_53; } else { -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_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_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); -return x_59; +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; } } } case 8: { -lean_object* x_60; lean_object* x_61; -x_60 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_61 = l_Lean_Compiler_ToLCNF_toLCNF_visitLet(x_1, x_60, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_61) == 0) +x_62 = l_Lean_Compiler_ToLCNF_toLCNF_visitLet(x_2, x_61, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_61); -x_64 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_62, x_3, x_4, x_5, x_6, x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_63, x_4, x_5, x_6, x_7, x_64); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_64; +return x_65; } else { -uint8_t x_65; +uint8_t x_66; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_65 = !lean_is_exclusive(x_61); -if (x_65 == 0) +x_66 = !lean_is_exclusive(x_62); +if (x_66 == 0) { -return x_61; +return x_62; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_61, 0); -x_67 = lean_ctor_get(x_61, 1); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_62, 0); +x_68 = lean_ctor_get(x_62, 1); +lean_inc(x_68); lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_61); -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; +lean_dec(x_62); +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; } } } case 9: { -lean_object* x_69; +lean_object* x_70; +lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_69 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_69) == 0) +x_70 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); -lean_dec(x_69); -x_72 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_70, x_3, x_4, x_5, x_6, x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_71, x_4, x_5, x_6, x_7, x_72); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_72; +return x_73; } else { -uint8_t x_73; +uint8_t x_74; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_73 = !lean_is_exclusive(x_69); -if (x_73 == 0) +x_74 = !lean_is_exclusive(x_70); +if (x_74 == 0) { -return x_69; +return x_70; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_69, 0); -x_75 = lean_ctor_get(x_69, 1); +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_70, 0); +x_76 = lean_ctor_get(x_70, 1); +lean_inc(x_76); lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_69); -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_dec(x_70); +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; } } } case 10: { -lean_object* x_77; lean_object* x_78; -x_77 = lean_ctor_get(x_1, 1); -lean_inc(x_77); +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_2, 1); +lean_inc(x_78); +lean_dec(x_2); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_78 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_77, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_78) == 0) +x_79 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_78, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_79, 0); lean_inc(x_80); -lean_dec(x_78); -x_81 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_79, x_3, x_4, x_5, x_6, x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_82 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_80, x_4, x_5, x_6, x_7, x_81); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_81; +return x_82; } else { -uint8_t x_82; +uint8_t x_83; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_82 = !lean_is_exclusive(x_78); -if (x_82 == 0) +x_83 = !lean_is_exclusive(x_79); +if (x_83 == 0) { -return x_78; +return x_79; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_78, 0); -x_84 = lean_ctor_get(x_78, 1); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_79, 0); +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_78); -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; +lean_dec(x_79); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } case 11: { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = lean_ctor_get(x_1, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_1, 1); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = lean_ctor_get(x_2, 0); lean_inc(x_87); -x_88 = lean_ctor_get(x_1, 2); +x_88 = lean_ctor_get(x_2, 1); lean_inc(x_88); +x_89 = lean_ctor_get(x_2, 2); +lean_inc(x_89); +lean_dec(x_2); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_89 = l_Lean_Compiler_ToLCNF_toLCNF_visitProj(x_86, x_87, x_88, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_89) == 0) +x_90 = l_Lean_Compiler_ToLCNF_toLCNF_visitProj(x_87, x_88, x_89, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_90) == 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_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_90, 0); lean_inc(x_91); -lean_dec(x_89); -x_92 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_90, x_3, x_4, x_5, x_6, x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_91, x_4, x_5, x_6, x_7, x_92); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_92; +return x_93; } else { -uint8_t x_93; +uint8_t x_94; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_93 = !lean_is_exclusive(x_89); -if (x_93 == 0) +x_94 = !lean_is_exclusive(x_90); +if (x_94 == 0) { -return x_89; +return x_90; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_89, 0); -x_95 = lean_ctor_get(x_89, 1); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_90, 0); +x_96 = lean_ctor_get(x_90, 1); +lean_inc(x_96); lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_89); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -return x_96; +lean_dec(x_90); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; } } } default: { -lean_object* x_97; -lean_inc(x_1); -x_97 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_1, x_3, x_4, x_5, x_6, x_7); +lean_object* x_98; +x_98 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__1(x_1, x_2, 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); -return x_97; +return x_98; } } } @@ -4664,7 +6400,7 @@ uint8_t x_19; x_19 = !lean_is_exclusive(x_4); 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; uint8_t x_36; +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; x_20 = lean_ctor_get(x_4, 10); lean_dec(x_20); x_21 = lean_ctor_get(x_4, 9); @@ -4691,154 +6427,167 @@ x_31 = lean_unsigned_to_nat(1u); x_32 = lean_nat_add(x_10, x_31); lean_dec(x_10); lean_ctor_set(x_4, 3, x_32); -x_33 = lean_st_ref_get(x_5, x_6); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_st_ref_get(x_3, x_34); -x_36 = !lean_is_exclusive(x_35); -if (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_35, 0); -x_38 = lean_ctor_get(x_35, 1); -x_39 = lean_ctor_get(x_37, 4); -lean_inc(x_39); -lean_dec(x_37); +x_33 = lean_box(x_2); lean_inc(x_1); -x_40 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_39, x_1); -if (lean_obj_tag(x_40) == 0) +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_1); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_st_ref_get(x_5, x_6); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_get(x_3, x_36); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_41; lean_object* x_42; -lean_free_object(x_35); -x_41 = lean_box(0); -x_42 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_1, x_41, x_2, x_3, x_4, x_5, x_38); -return x_42; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_37, 0); +x_40 = lean_ctor_get(x_37, 1); +x_41 = lean_ctor_get(x_39, 4); +lean_inc(x_41); +lean_dec(x_39); +lean_inc(x_34); +x_42 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_41, x_34); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_free_object(x_37); +x_43 = lean_box(0); +x_44 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_34, x_1, x_43, x_2, x_3, x_4, x_5, x_40); +return x_44; } else { -lean_object* x_43; +lean_object* x_45; +lean_dec(x_34); lean_dec(x_4); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -lean_dec(x_40); -lean_ctor_set(x_35, 0, x_43); -return x_35; +x_45 = lean_ctor_get(x_42, 0); +lean_inc(x_45); +lean_dec(x_42); +lean_ctor_set(x_37, 0, x_45); +return x_37; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_ctor_get(x_35, 0); -x_45 = lean_ctor_get(x_35, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_35); -x_46 = lean_ctor_get(x_44, 4); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_37, 0); +x_47 = lean_ctor_get(x_37, 1); +lean_inc(x_47); lean_inc(x_46); -lean_dec(x_44); -lean_inc(x_1); -x_47 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_46, x_1); -if (lean_obj_tag(x_47) == 0) +lean_dec(x_37); +x_48 = lean_ctor_get(x_46, 4); +lean_inc(x_48); +lean_dec(x_46); +lean_inc(x_34); +x_49 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_48, x_34); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_48; lean_object* x_49; -x_48 = lean_box(0); -x_49 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_1, x_48, x_2, x_3, x_4, x_5, x_45); -return x_49; +lean_object* x_50; lean_object* x_51; +x_50 = lean_box(0); +x_51 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_34, x_1, x_50, x_2, x_3, x_4, x_5, x_47); +return x_51; } else { -lean_object* x_50; lean_object* x_51; +lean_object* x_52; lean_object* x_53; +lean_dec(x_34); lean_dec(x_4); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_50 = lean_ctor_get(x_47, 0); -lean_inc(x_50); -lean_dec(x_47); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_45); -return x_51; +x_52 = lean_ctor_get(x_49, 0); +lean_inc(x_52); +lean_dec(x_49); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_47); +return x_53; } } } 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; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +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_dec(x_4); -x_52 = lean_unsigned_to_nat(1u); -x_53 = lean_nat_add(x_10, x_52); +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_add(x_10, x_54); lean_dec(x_10); -x_54 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_54, 0, x_7); -lean_ctor_set(x_54, 1, x_8); -lean_ctor_set(x_54, 2, x_9); -lean_ctor_set(x_54, 3, x_53); -lean_ctor_set(x_54, 4, x_11); -lean_ctor_set(x_54, 5, x_12); -lean_ctor_set(x_54, 6, x_13); -lean_ctor_set(x_54, 7, x_14); -lean_ctor_set(x_54, 8, x_15); -lean_ctor_set(x_54, 9, x_16); -lean_ctor_set(x_54, 10, x_17); -x_55 = lean_st_ref_get(x_5, x_6); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = lean_st_ref_get(x_3, 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); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_60 = x_57; +x_56 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_56, 0, x_7); +lean_ctor_set(x_56, 1, x_8); +lean_ctor_set(x_56, 2, x_9); +lean_ctor_set(x_56, 3, x_55); +lean_ctor_set(x_56, 4, x_11); +lean_ctor_set(x_56, 5, x_12); +lean_ctor_set(x_56, 6, x_13); +lean_ctor_set(x_56, 7, x_14); +lean_ctor_set(x_56, 8, x_15); +lean_ctor_set(x_56, 9, x_16); +lean_ctor_set(x_56, 10, x_17); +x_57 = lean_box(x_2); +lean_inc(x_1); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_1); +lean_ctor_set(x_58, 1, x_57); +x_59 = lean_st_ref_get(x_5, x_6); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = lean_st_ref_get(x_3, x_60); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_64 = x_61; } else { - lean_dec_ref(x_57); - x_60 = lean_box(0); + lean_dec_ref(x_61); + x_64 = lean_box(0); } -x_61 = lean_ctor_get(x_58, 4); -lean_inc(x_61); -lean_dec(x_58); -lean_inc(x_1); -x_62 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_61, x_1); -if (lean_obj_tag(x_62) == 0) +x_65 = lean_ctor_get(x_62, 4); +lean_inc(x_65); +lean_dec(x_62); +lean_inc(x_58); +x_66 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__7(x_65, x_58); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_63; lean_object* x_64; -lean_dec(x_60); -x_63 = lean_box(0); -x_64 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_1, x_63, x_2, x_3, x_54, x_5, x_59); -return x_64; +lean_object* x_67; lean_object* x_68; +lean_dec(x_64); +x_67 = lean_box(0); +x_68 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_58, x_1, x_67, x_2, x_3, x_56, x_5, x_63); +return x_68; } else { -lean_object* x_65; lean_object* x_66; -lean_dec(x_54); +lean_object* x_69; lean_object* x_70; +lean_dec(x_58); +lean_dec(x_56); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_65 = lean_ctor_get(x_62, 0); -lean_inc(x_65); -lean_dec(x_62); -if (lean_is_scalar(x_60)) { - x_66 = lean_alloc_ctor(0, 2, 0); +x_69 = lean_ctor_get(x_66, 0); +lean_inc(x_69); +lean_dec(x_66); +if (lean_is_scalar(x_64)) { + x_70 = lean_alloc_ctor(0, 2, 0); } else { - x_66 = x_60; + x_70 = x_64; } -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_59); -return x_66; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_63); +return x_70; } } } else { -lean_object* x_67; +lean_object* x_71; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -4850,11 +6599,11 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_1); -x_67 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__10(x_12, x_2, x_3, x_4, x_5, x_6); +x_71 = l_Lean_throwMaxRecDepthAt___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__10(x_12, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_67; +return x_71; } } } @@ -4863,7 +6612,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitLet(lean_object* x_1 { if (lean_obj_tag(x_1) == 8) { -lean_object* x_8; 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_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_8; 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_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_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_1, 1); @@ -4877,192 +6626,192 @@ x_12 = lean_expr_instantiate_rev(x_9, x_2); lean_dec(x_9); x_13 = lean_expr_instantiate_rev(x_10, x_2); lean_dec(x_10); -x_67 = lean_st_ref_get(x_6, x_7); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = lean_st_ref_get(x_4, x_68); -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); -lean_inc(x_71); -lean_dec(x_69); -x_72 = lean_ctor_get(x_70, 0); -lean_inc(x_72); -lean_dec(x_70); -x_73 = lean_box(0); -x_74 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; -x_75 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; -x_76 = lean_unsigned_to_nat(0u); -x_77 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_72); -lean_ctor_set(x_77, 2, x_75); -lean_ctor_set(x_77, 3, x_73); -lean_ctor_set(x_77, 4, x_76); -lean_ctor_set(x_77, 5, x_73); -x_78 = lean_st_ref_get(x_6, x_71); -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_80 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; -x_81 = lean_st_mk_ref(x_80, x_79); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); +x_46 = lean_st_ref_get(x_6, x_7); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_st_ref_get(x_4, x_47); +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 = lean_ctor_get(x_49, 0); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_box(0); +x_53 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; +x_54 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +x_55 = lean_unsigned_to_nat(0u); +x_56 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_51); +lean_ctor_set(x_56, 2, x_54); +lean_ctor_set(x_56, 3, x_52); +lean_ctor_set(x_56, 4, x_55); +lean_ctor_set(x_56, 5, x_52); +x_57 = lean_st_ref_get(x_6, x_50); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; +x_60 = lean_st_mk_ref(x_59, x_58); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_82); -lean_inc(x_77); +lean_inc(x_61); +lean_inc(x_56); lean_inc(x_12); -x_84 = l_Lean_Meta_isProp(x_12, x_77, x_82, x_5, x_6, x_83); -if (lean_obj_tag(x_84) == 0) +x_63 = l_Lean_Meta_isProp(x_12, x_56, x_61, x_5, x_6, x_62); +if (lean_obj_tag(x_63) == 0) { -lean_object* x_85; uint8_t x_86; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_unbox(x_85); -if (x_86 == 0) +lean_object* x_64; uint8_t x_65; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_unbox(x_64); +if (x_65 == 0) { -lean_object* x_87; lean_object* x_88; -lean_dec(x_85); -x_87 = lean_ctor_get(x_84, 1); -lean_inc(x_87); -lean_dec(x_84); +lean_object* x_66; lean_object* x_67; +lean_dec(x_64); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_82); +lean_inc(x_61); lean_inc(x_12); -x_88 = l_Lean_Meta_isTypeFormerType(x_12, x_77, x_82, x_5, x_6, x_87); -if (lean_obj_tag(x_88) == 0) +x_67 = l_Lean_Meta_isTypeFormerType(x_12, x_56, x_61, x_5, x_6, x_66); +if (lean_obj_tag(x_67) == 0) { -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_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_st_ref_get(x_6, x_90); -x_92 = lean_ctor_get(x_91, 1); -lean_inc(x_92); -lean_dec(x_91); -x_93 = lean_st_ref_get(x_82, x_92); -lean_dec(x_82); -x_94 = !lean_is_exclusive(x_93); -if (x_94 == 0) +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_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_st_ref_get(x_6, x_69); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = lean_st_ref_get(x_61, x_71); +lean_dec(x_61); +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) { -lean_object* x_95; -x_95 = lean_ctor_get(x_93, 0); -lean_dec(x_95); -lean_ctor_set(x_93, 0, x_89); -x_14 = x_93; -goto block_66; +lean_object* x_74; +x_74 = lean_ctor_get(x_72, 0); +lean_dec(x_74); +lean_ctor_set(x_72, 0, x_68); +x_14 = x_72; +goto block_45; } else { -lean_object* x_96; lean_object* x_97; -x_96 = lean_ctor_get(x_93, 1); -lean_inc(x_96); -lean_dec(x_93); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_89); -lean_ctor_set(x_97, 1, x_96); -x_14 = x_97; -goto block_66; +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +lean_dec(x_72); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_68); +lean_ctor_set(x_76, 1, x_75); +x_14 = x_76; +goto block_45; } } else { -uint8_t x_98; -lean_dec(x_82); -x_98 = !lean_is_exclusive(x_88); -if (x_98 == 0) +uint8_t x_77; +lean_dec(x_61); +x_77 = !lean_is_exclusive(x_67); +if (x_77 == 0) { -x_14 = x_88; -goto block_66; +x_14 = x_67; +goto block_45; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_88, 0); -x_100 = lean_ctor_get(x_88, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_88); -x_101 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -x_14 = x_101; -goto block_66; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_67, 0); +x_79 = lean_ctor_get(x_67, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_67); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_14 = x_80; +goto block_45; } } } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -lean_dec(x_77); -x_102 = lean_ctor_get(x_84, 1); -lean_inc(x_102); -lean_dec(x_84); -x_103 = lean_st_ref_get(x_6, x_102); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -lean_dec(x_103); -x_105 = lean_st_ref_get(x_82, x_104); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +lean_dec(x_56); +x_81 = lean_ctor_get(x_63, 1); +lean_inc(x_81); +lean_dec(x_63); +x_82 = lean_st_ref_get(x_6, x_81); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); lean_dec(x_82); -x_106 = !lean_is_exclusive(x_105); -if (x_106 == 0) +x_84 = lean_st_ref_get(x_61, x_83); +lean_dec(x_61); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_107; -x_107 = lean_ctor_get(x_105, 0); -lean_dec(x_107); -lean_ctor_set(x_105, 0, x_85); -x_14 = x_105; -goto block_66; +lean_object* x_86; +x_86 = lean_ctor_get(x_84, 0); +lean_dec(x_86); +lean_ctor_set(x_84, 0, x_64); +x_14 = x_84; +goto block_45; } else { -lean_object* x_108; lean_object* x_109; -x_108 = lean_ctor_get(x_105, 1); -lean_inc(x_108); -lean_dec(x_105); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_85); -lean_ctor_set(x_109, 1, x_108); -x_14 = x_109; -goto block_66; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_84, 1); +lean_inc(x_87); +lean_dec(x_84); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_64); +lean_ctor_set(x_88, 1, x_87); +x_14 = x_88; +goto block_45; } } } else { -uint8_t x_110; -lean_dec(x_82); -lean_dec(x_77); -x_110 = !lean_is_exclusive(x_84); -if (x_110 == 0) +uint8_t x_89; +lean_dec(x_61); +lean_dec(x_56); +x_89 = !lean_is_exclusive(x_63); +if (x_89 == 0) { -x_14 = x_84; -goto block_66; +x_14 = x_63; +goto block_45; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_84, 0); -x_112 = lean_ctor_get(x_84, 1); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_84); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -x_14 = x_113; -goto block_66; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_63, 0); +x_91 = lean_ctor_get(x_63, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_63); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +x_14 = x_92; +goto block_45; } } -block_66: +block_45: { if (lean_obj_tag(x_14) == 0) { @@ -5073,100 +6822,55 @@ x_16 = lean_unbox(x_15); lean_dec(x_15); 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; 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_17; lean_object* x_18; x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); lean_dec(x_14); -x_18 = lean_st_ref_get(x_6, x_17); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_st_ref_get(x_4, 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_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_box(0); -x_25 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1; -x_26 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_28, 0, x_25); -lean_ctor_set(x_28, 1, x_23); -lean_ctor_set(x_28, 2, x_26); -lean_ctor_set(x_28, 3, x_24); -lean_ctor_set(x_28, 4, x_27); -lean_ctor_set(x_28, 5, x_24); -x_29 = lean_st_ref_get(x_6, x_22); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13; -x_32 = lean_st_mk_ref(x_31, x_30); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_33); lean_inc(x_12); -x_35 = l_Lean_Compiler_toLCNFType(x_12, x_28, x_33, x_5, x_6, x_34); -if (lean_obj_tag(x_35) == 0) +x_18 = l_Lean_Compiler_ToLCNF_toLCNFType(x_12, x_3, x_4, x_5, x_6, x_17); +if (lean_obj_tag(x_18) == 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; lean_object* x_43; -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_st_ref_get(x_6, x_37); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -lean_dec(x_38); -x_40 = lean_st_ref_get(x_33, x_39); -lean_dec(x_33); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = 1; +lean_object* x_19; lean_object* x_20; uint8_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 = 1; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_13); -x_43 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_13, x_42, x_4, x_5, x_6, x_41); -if (lean_obj_tag(x_43) == 0) +x_22 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_13, x_21, x_4, x_5, x_6, x_20); +if (lean_obj_tag(x_22) == 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_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); +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_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); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_46 = l_Lean_Compiler_ToLCNF_mkLetDecl(x_8, x_12, x_13, x_36, x_44, x_3, x_4, x_5, x_6, x_45); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = lean_array_push(x_2, x_47); +x_25 = l_Lean_Compiler_ToLCNF_mkLetDecl(x_8, x_12, x_13, x_19, x_23, x_3, x_4, x_5, x_6, x_24); +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_array_push(x_2, x_26); x_1 = x_11; -x_2 = x_49; -x_7 = x_48; +x_2 = x_28; +x_7 = x_27; goto _start; } else { -uint8_t x_51; -lean_dec(x_36); +uint8_t x_30; +lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5175,30 +6879,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_51 = !lean_is_exclusive(x_43); -if (x_51 == 0) +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) { -return x_43; +return x_22; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_43, 0); -x_53 = lean_ctor_get(x_43, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_43); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +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; } } } else { -uint8_t x_55; -lean_dec(x_33); +uint8_t x_34; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5207,44 +6910,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_55 = !lean_is_exclusive(x_35); -if (x_55 == 0) +x_34 = !lean_is_exclusive(x_18); +if (x_34 == 0) { -return x_35; +return x_18; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_35, 0); -x_57 = lean_ctor_get(x_35, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_35); -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_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_18, 0); +x_36 = lean_ctor_get(x_18, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_18); +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_59; lean_object* x_60; +lean_object* x_38; lean_object* x_39; lean_dec(x_12); lean_dec(x_8); -x_59 = lean_ctor_get(x_14, 1); -lean_inc(x_59); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); lean_dec(x_14); -x_60 = lean_array_push(x_2, x_13); +x_39 = lean_array_push(x_2, x_13); x_1 = x_11; -x_2 = x_60; -x_7 = x_59; +x_2 = x_39; +x_7 = x_38; goto _start; } } else { -uint8_t x_62; +uint8_t x_41; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5253,35 +6956,35 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_14); -if (x_62 == 0) +x_41 = !lean_is_exclusive(x_14); +if (x_41 == 0) { return x_14; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_14, 0); -x_64 = lean_ctor_get(x_14, 1); -lean_inc(x_64); -lean_inc(x_63); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_14, 0); +x_43 = lean_ctor_get(x_14, 1); +lean_inc(x_43); +lean_inc(x_42); lean_dec(x_14); -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; +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } } else { -lean_object* x_114; lean_object* x_115; -x_114 = lean_expr_instantiate_rev(x_1, x_2); +lean_object* x_93; lean_object* x_94; +x_93 = lean_expr_instantiate_rev(x_1, x_2); lean_dec(x_2); lean_dec(x_1); -x_115 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_114, x_3, x_4, x_5, x_6, x_7); -return x_115; +x_94 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_93, x_3, x_4, x_5, x_6, x_7); +return x_94; } } } @@ -7002,9 +8705,7 @@ x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); x_19 = lean_unsigned_to_nat(0u); -lean_inc_n(x_3, 2); -lean_inc(x_2); -lean_inc(x_1); +lean_inc(x_3); x_20 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_20, 0, x_1); lean_ctor_set(x_20, 1, x_18); @@ -7016,7 +8717,6 @@ x_21 = lean_st_ref_get(x_11, x_17); x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); lean_dec(x_21); -lean_inc(x_4); x_23 = lean_st_mk_ref(x_4, x_22); x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); @@ -7046,10 +8746,6 @@ lean_dec(x_27); if (x_32 == 0) { lean_object* x_33; uint8_t x_34; lean_object* x_35; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); lean_dec(x_31); @@ -7059,136 +8755,40 @@ 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_36; lean_object* x_37; x_36 = lean_ctor_get(x_31, 1); lean_inc(x_36); lean_dec(x_31); -x_37 = lean_st_ref_get(x_11, x_36); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = lean_st_ref_get(x_9, x_38); +x_37 = l_Lean_Compiler_ToLCNF_toLCNFType(x_6, x_8, x_9, x_10, x_11, x_36); lean_dec(x_9); -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_ctor_get(x_40, 0); -lean_inc(x_42); -lean_dec(x_40); -lean_inc(x_3); -x_43 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_2); -lean_ctor_set(x_43, 3, x_3); -lean_ctor_set(x_43, 4, x_19); -lean_ctor_set(x_43, 5, x_3); -x_44 = lean_st_ref_get(x_11, x_41); -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -lean_dec(x_44); -x_46 = lean_st_mk_ref(x_4, x_45); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -lean_inc(x_11); -lean_inc(x_47); -x_49 = l_Lean_Compiler_toLCNFType(x_6, x_43, x_47, x_10, x_11, x_48); -if (lean_obj_tag(x_49) == 0) -{ -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_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_st_ref_get(x_11, x_51); -lean_dec(x_11); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = lean_st_ref_get(x_47, x_53); -lean_dec(x_47); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) -{ -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -lean_ctor_set(x_54, 0, x_50); -return x_54; -} -else -{ -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_50); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} -} -else -{ -uint8_t x_59; -lean_dec(x_47); -lean_dec(x_11); -x_59 = !lean_is_exclusive(x_49); -if (x_59 == 0) -{ -return x_49; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_49, 0); -x_61 = lean_ctor_get(x_49, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_49); -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; -} -} +return x_37; } } else { -uint8_t x_63; +uint8_t x_38; lean_dec(x_24); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_63 = !lean_is_exclusive(x_26); -if (x_63 == 0) +x_38 = !lean_is_exclusive(x_26); +if (x_38 == 0) { return x_26; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_26, 0); -x_65 = lean_ctor_get(x_26, 1); -lean_inc(x_65); -lean_inc(x_64); +lean_object* x_39; lean_object* x_40; lean_object* 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_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; +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; } } } @@ -7611,7 +9211,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___closed__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; x_1 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitProjFn___closed__1; -x_3 = lean_unsigned_to_nat(393u); +x_3 = lean_unsigned_to_nat(404u); x_4 = lean_unsigned_to_nat(45u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8563,7 +10163,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___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_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___closed__1; -x_3 = lean_unsigned_to_nat(350u); +x_3 = lean_unsigned_to_nat(361u); x_4 = lean_unsigned_to_nat(42u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8576,7 +10176,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___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_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___closed__1; -x_3 = lean_unsigned_to_nat(352u); +x_3 = lean_unsigned_to_nat(363u); x_4 = lean_unsigned_to_nat(56u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -10208,441 +11808,434 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_Compiler_ToLCNF_toLCNF_visitCases___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) { _start: { -lean_object* x_7; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); -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_object* x_11; lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_Compiler_toLCNFType(x_8, x_2, x_3, x_4, x_5, x_9); -return x_10; -} -else -{ -uint8_t x_11; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_11 = !lean_is_exclusive(x_7); -if (x_11 == 0) -{ -return x_7; -} -else +lean_inc(x_8); +x_11 = l_Lean_Compiler_ToLCNF_toLCNFType(x_5, x_6, x_7, x_8, x_9, 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_7, 0); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); +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_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_7); -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_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__2(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) { -_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; lean_object* x_23; -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -lean_inc(x_11); -x_14 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__1(x_11, x_12, x_11, x_13, x_2, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); lean_dec(x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_unsigned_to_nat(1u); +lean_inc(x_14); +x_17 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__1(x_14, x_15, x_14, x_16, x_2, x_6, x_7, x_8, x_9, x_13); lean_dec(x_14); -x_17 = lean_ctor_get(x_1, 2); -lean_inc(x_17); -x_18 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_17, 2); -lean_inc(x_22); +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_9); -lean_inc(x_8); -lean_inc(x_7); +x_20 = lean_ctor_get(x_1, 2); lean_inc(x_20); -x_23 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__2(x_20, x_21, x_20, x_22, x_19, x_6, x_7, x_8, x_9, x_16); -lean_dec(x_22); -lean_dec(x_20); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = lean_ctor_get(x_23, 0); +x_21 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +x_25 = lean_ctor_get(x_20, 2); lean_inc(x_25); +lean_dec(x_20); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_23); +x_26 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__2(x_23, x_24, x_23, x_25, x_22, x_6, x_7, x_8, x_9, x_19); +lean_dec(x_25); lean_dec(x_23); -x_26 = !lean_is_exclusive(x_24); -if (x_26 == 0) +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; lean_object* x_35; lean_object* x_36; -x_27 = lean_ctor_get(x_24, 1); -x_28 = lean_ctor_get(x_1, 4); +lean_object* x_27; lean_object* x_28; uint8_t 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); -x_29 = lean_array_get_size(x_28); -x_30 = l_Array_toSubarray___rarg(x_28, x_12, x_29); -x_31 = lean_ctor_get(x_1, 3); +lean_dec(x_26); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +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_30 = lean_ctor_get(x_27, 1); +x_31 = lean_ctor_get(x_1, 4); lean_inc(x_31); -lean_ctor_set(x_24, 1, x_5); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_24); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 0); +x_32 = lean_array_get_size(x_31); +x_33 = l_Array_toSubarray___rarg(x_31, x_15, x_32); +x_34 = lean_ctor_get(x_1, 3); lean_inc(x_34); -x_35 = lean_ctor_get(x_31, 2); -lean_inc(x_35); -lean_dec(x_31); +lean_ctor_set(x_27, 1, x_12); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_27); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_34, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_34, 2); +lean_inc(x_38); +lean_dec(x_34); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_33); -x_36 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3(x_33, x_34, x_33, x_35, x_32, x_6, x_7, x_8, x_9, x_25); -lean_dec(x_35); -lean_dec(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; lean_object* x_41; lean_object* x_42; uint8_t x_43; -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_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); +lean_inc(x_36); +x_39 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3(x_36, x_37, x_36, x_38, x_35, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_38); lean_dec(x_36); -x_40 = lean_ctor_get(x_38, 0); +if (lean_obj_tag(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; uint8_t x_46; +x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); +x_41 = lean_ctor_get(x_40, 1); lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_array_get_size(x_27); -x_43 = lean_nat_dec_lt(x_12, 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; uint8_t x_50; -lean_dec(x_42); -lean_dec(x_27); -x_44 = lean_ctor_get(x_1, 5); +lean_dec(x_40); +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_41, 1); lean_inc(x_44); +lean_dec(x_41); +x_45 = lean_array_get_size(x_30); +x_46 = lean_nat_dec_lt(x_15, x_45); +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_object* x_52; uint8_t x_53; +lean_dec(x_45); +lean_dec(x_30); +x_47 = lean_ctor_get(x_1, 5); +lean_inc(x_47); lean_dec(x_1); -x_45 = lean_array_set(x_40, x_44, x_41); -lean_dec(x_44); +x_48 = lean_array_set(x_43, x_47, x_44); +lean_dec(x_47); lean_inc(x_3); -lean_inc(x_45); -x_46 = l_Array_toSubarray___rarg(x_45, x_12, x_3); -x_47 = l_Array_ofSubarray___rarg(x_46); -x_48 = l_Lean_mkAppN(x_4, x_47); -x_49 = lean_array_get_size(x_45); -x_50 = lean_nat_dec_eq(x_49, x_3); -lean_dec(x_49); -if (x_50 == 0) +lean_inc(x_48); +x_49 = l_Array_toSubarray___rarg(x_48, x_15, x_3); +x_50 = l_Array_ofSubarray___rarg(x_49); +x_51 = l_Lean_mkAppN(x_4, x_50); +x_52 = lean_array_get_size(x_48); +x_53 = lean_nat_dec_eq(x_52, x_3); +lean_dec(x_52); +if (x_53 == 0) { -lean_object* x_51; -x_51 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_48, x_45, x_3, x_6, x_7, x_8, x_9, x_39); -return x_51; +lean_object* x_54; +x_54 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_51, x_48, x_3, x_6, x_7, x_8, x_9, x_42); +return x_54; } else { -lean_object* x_52; -lean_dec(x_45); +lean_object* x_55; +lean_dec(x_48); lean_dec(x_3); -x_52 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_48, x_6, x_7, x_8, x_9, x_39); +x_55 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_51, x_6, x_7, x_8, x_9, x_42); lean_dec(x_7); -return x_52; +return x_55; } } else { -size_t x_53; size_t 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; uint8_t x_64; -x_53 = lean_usize_of_nat(x_42); -lean_dec(x_42); -x_54 = 0; -x_55 = l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4(x_27, x_53, x_54, x_41, x_6, x_7, x_8, x_9, x_39); -lean_dec(x_27); -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_1, 5); -lean_inc(x_58); -lean_dec(x_1); -x_59 = lean_array_set(x_40, x_58, x_56); +size_t x_56; size_t 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_56 = lean_usize_of_nat(x_45); +lean_dec(x_45); +x_57 = 0; +x_58 = l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4(x_30, x_56, x_57, x_44, x_6, x_7, x_8, x_9, x_42); +lean_dec(x_30); +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_1, 5); +lean_inc(x_61); +lean_dec(x_1); +x_62 = lean_array_set(x_43, x_61, x_59); +lean_dec(x_61); lean_inc(x_3); -lean_inc(x_59); -x_60 = l_Array_toSubarray___rarg(x_59, x_12, x_3); -x_61 = l_Array_ofSubarray___rarg(x_60); -x_62 = l_Lean_mkAppN(x_4, x_61); -x_63 = lean_array_get_size(x_59); -x_64 = lean_nat_dec_eq(x_63, x_3); -lean_dec(x_63); -if (x_64 == 0) +lean_inc(x_62); +x_63 = l_Array_toSubarray___rarg(x_62, x_15, x_3); +x_64 = l_Array_ofSubarray___rarg(x_63); +x_65 = l_Lean_mkAppN(x_4, x_64); +x_66 = lean_array_get_size(x_62); +x_67 = lean_nat_dec_eq(x_66, x_3); +lean_dec(x_66); +if (x_67 == 0) { -lean_object* x_65; -x_65 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_62, x_59, x_3, x_6, x_7, x_8, x_9, x_57); -return x_65; +lean_object* x_68; +x_68 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_65, x_62, x_3, x_6, x_7, x_8, x_9, x_60); +return x_68; } else { -lean_object* x_66; -lean_dec(x_59); +lean_object* x_69; +lean_dec(x_62); lean_dec(x_3); -x_66 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_62, x_6, x_7, x_8, x_9, x_57); +x_69 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_65, x_6, x_7, x_8, x_9, x_60); lean_dec(x_7); -return x_66; +return x_69; } } } else { -uint8_t x_67; -lean_dec(x_27); +uint8_t x_70; +lean_dec(x_30); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_67 = !lean_is_exclusive(x_36); -if (x_67 == 0) +x_70 = !lean_is_exclusive(x_39); +if (x_70 == 0) { -return x_36; +return x_39; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_36, 0); -x_69 = lean_ctor_get(x_36, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_36); -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_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_39, 0); +x_72 = lean_ctor_get(x_39, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_39); +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_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_71 = lean_ctor_get(x_24, 0); -x_72 = lean_ctor_get(x_24, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_24); -x_73 = lean_ctor_get(x_1, 4); -lean_inc(x_73); -x_74 = lean_array_get_size(x_73); -x_75 = l_Array_toSubarray___rarg(x_73, x_12, x_74); -x_76 = lean_ctor_get(x_1, 3); +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_74 = lean_ctor_get(x_27, 0); +x_75 = lean_ctor_get(x_27, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_27); +x_76 = lean_ctor_get(x_1, 4); lean_inc(x_76); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_71); -lean_ctor_set(x_77, 1, x_5); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_77); -x_79 = lean_ctor_get(x_76, 1); +x_77 = lean_array_get_size(x_76); +x_78 = l_Array_toSubarray___rarg(x_76, x_15, x_77); +x_79 = lean_ctor_get(x_1, 3); lean_inc(x_79); -x_80 = lean_ctor_get(x_76, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_76, 2); -lean_inc(x_81); -lean_dec(x_76); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_74); +lean_ctor_set(x_80, 1, x_12); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_ctor_get(x_79, 1); +lean_inc(x_82); +x_83 = lean_ctor_get(x_79, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_79, 2); +lean_inc(x_84); +lean_dec(x_79); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_79); -x_82 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3(x_79, x_80, x_79, x_81, x_78, x_6, x_7, x_8, x_9, x_25); -lean_dec(x_81); -lean_dec(x_79); -if (lean_obj_tag(x_82) == 0) -{ -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; -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -lean_dec(x_83); -x_85 = lean_ctor_get(x_82, 1); -lean_inc(x_85); +lean_inc(x_82); +x_85 = l_Std_Range_forIn_loop___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__3(x_82, x_83, x_82, x_84, x_81, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_84); lean_dec(x_82); -x_86 = lean_ctor_get(x_84, 0); +if (lean_obj_tag(x_85) == 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; uint8_t x_92; +x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); -x_87 = lean_ctor_get(x_84, 1); +x_87 = lean_ctor_get(x_86, 1); lean_inc(x_87); -lean_dec(x_84); -x_88 = lean_array_get_size(x_72); -x_89 = lean_nat_dec_lt(x_12, 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; uint8_t x_96; -lean_dec(x_88); -lean_dec(x_72); -x_90 = lean_ctor_get(x_1, 5); +lean_dec(x_86); +x_88 = lean_ctor_get(x_85, 1); +lean_inc(x_88); +lean_dec(x_85); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_87, 1); lean_inc(x_90); +lean_dec(x_87); +x_91 = lean_array_get_size(x_75); +x_92 = lean_nat_dec_lt(x_15, 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; uint8_t x_99; +lean_dec(x_91); +lean_dec(x_75); +x_93 = lean_ctor_get(x_1, 5); +lean_inc(x_93); lean_dec(x_1); -x_91 = lean_array_set(x_86, x_90, x_87); -lean_dec(x_90); +x_94 = lean_array_set(x_89, x_93, x_90); +lean_dec(x_93); lean_inc(x_3); -lean_inc(x_91); -x_92 = l_Array_toSubarray___rarg(x_91, x_12, x_3); -x_93 = l_Array_ofSubarray___rarg(x_92); -x_94 = l_Lean_mkAppN(x_4, x_93); -x_95 = lean_array_get_size(x_91); -x_96 = lean_nat_dec_eq(x_95, x_3); -lean_dec(x_95); -if (x_96 == 0) +lean_inc(x_94); +x_95 = l_Array_toSubarray___rarg(x_94, x_15, x_3); +x_96 = l_Array_ofSubarray___rarg(x_95); +x_97 = l_Lean_mkAppN(x_4, x_96); +x_98 = lean_array_get_size(x_94); +x_99 = lean_nat_dec_eq(x_98, x_3); +lean_dec(x_98); +if (x_99 == 0) { -lean_object* x_97; -x_97 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_94, x_91, x_3, x_6, x_7, x_8, x_9, x_85); -return x_97; +lean_object* x_100; +x_100 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_97, x_94, x_3, x_6, x_7, x_8, x_9, x_88); +return x_100; } else { -lean_object* x_98; -lean_dec(x_91); +lean_object* x_101; +lean_dec(x_94); lean_dec(x_3); -x_98 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_94, x_6, x_7, x_8, x_9, x_85); +x_101 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_97, x_6, x_7, x_8, x_9, x_88); lean_dec(x_7); -return x_98; +return x_101; } } else { -size_t x_99; size_t 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; uint8_t x_110; -x_99 = lean_usize_of_nat(x_88); -lean_dec(x_88); -x_100 = 0; -x_101 = l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4(x_72, x_99, x_100, x_87, x_6, x_7, x_8, x_9, x_85); -lean_dec(x_72); -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_ctor_get(x_1, 5); -lean_inc(x_104); -lean_dec(x_1); -x_105 = lean_array_set(x_86, x_104, x_102); +size_t x_102; size_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; uint8_t x_113; +x_102 = lean_usize_of_nat(x_91); +lean_dec(x_91); +x_103 = 0; +x_104 = l_Array_foldrMUnsafe_fold___at_Lean_Compiler_ToLCNF_toLCNF_visitCases___spec__4(x_75, x_102, x_103, x_90, x_6, x_7, x_8, x_9, x_88); +lean_dec(x_75); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); lean_dec(x_104); +x_107 = lean_ctor_get(x_1, 5); +lean_inc(x_107); +lean_dec(x_1); +x_108 = lean_array_set(x_89, x_107, x_105); +lean_dec(x_107); lean_inc(x_3); -lean_inc(x_105); -x_106 = l_Array_toSubarray___rarg(x_105, x_12, x_3); -x_107 = l_Array_ofSubarray___rarg(x_106); -x_108 = l_Lean_mkAppN(x_4, x_107); -x_109 = lean_array_get_size(x_105); -x_110 = lean_nat_dec_eq(x_109, x_3); -lean_dec(x_109); -if (x_110 == 0) +lean_inc(x_108); +x_109 = l_Array_toSubarray___rarg(x_108, x_15, x_3); +x_110 = l_Array_ofSubarray___rarg(x_109); +x_111 = l_Lean_mkAppN(x_4, x_110); +x_112 = lean_array_get_size(x_108); +x_113 = lean_nat_dec_eq(x_112, x_3); +lean_dec(x_112); +if (x_113 == 0) { -lean_object* x_111; -x_111 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_108, x_105, x_3, x_6, x_7, x_8, x_9, x_103); -return x_111; +lean_object* x_114; +x_114 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_111, x_108, x_3, x_6, x_7, x_8, x_9, x_106); +return x_114; } else { -lean_object* x_112; -lean_dec(x_105); +lean_object* x_115; +lean_dec(x_108); lean_dec(x_3); -x_112 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_108, x_6, x_7, x_8, x_9, x_103); +x_115 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_111, x_6, x_7, x_8, x_9, x_106); lean_dec(x_7); -return x_112; +return x_115; } } } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_72); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_75); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_113 = lean_ctor_get(x_82, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_82, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_115 = x_82; +x_116 = lean_ctor_get(x_85, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_85, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_118 = x_85; } else { - lean_dec_ref(x_82); - x_115 = lean_box(0); + lean_dec_ref(x_85); + x_118 = lean_box(0); } -if (lean_is_scalar(x_115)) { - x_116 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); } else { - x_116 = x_115; + x_119 = x_118; } -lean_ctor_set(x_116, 0, x_113); -lean_ctor_set(x_116, 1, x_114); -return x_116; +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; } } } else { -uint8_t x_117; +uint8_t x_120; +lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_117 = !lean_is_exclusive(x_23); -if (x_117 == 0) +x_120 = !lean_is_exclusive(x_26); +if (x_120 == 0) { -return x_23; +return x_26; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_23, 0); -x_119 = lean_ctor_get(x_23, 1); -lean_inc(x_119); -lean_inc(x_118); -lean_dec(x_23); -x_120 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -return x_120; +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_26, 0); +x_122 = lean_ctor_get(x_26, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_26); +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_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_11); +if (x_124 == 0) +{ +return x_11; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_11, 0); +x_126 = lean_ctor_get(x_11, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_11); +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; } } } @@ -10670,12 +12263,12 @@ x_17 = l_Array_toSubarray___rarg(x_15, x_9, x_8); x_18 = l_Array_ofSubarray___rarg(x_17); lean_inc(x_16); x_19 = l_Lean_mkAppN(x_16, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1), 6, 1); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_inferType___boxed), 6, 1); lean_closure_set(x_20, 0, x_19); x_21 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___boxed), 6, 1); lean_closure_set(x_21, 0, x_20); lean_inc(x_8); -x_22 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__2___boxed), 10, 4); +x_22 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1___boxed), 10, 4); lean_closure_set(x_22, 0, x_1); lean_closure_set(x_22, 1, x_15); lean_closure_set(x_22, 2, x_8); @@ -11226,7 +12819,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitFalseRec___lambda__1 lean_object* x_7; lean_inc(x_5); lean_inc(x_4); -x_7 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Compiler_ToLCNF_toLCNFType(x_1, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_7) == 0) { lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -11235,31 +12828,67 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_8, x_2, x_3, x_4, x_5, x_9); +lean_inc(x_5); +lean_inc(x_4); +x_10 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___spec__1(x_8, x_2, x_3, x_4, x_5, 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 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_11, x_2, x_3, x_4, x_5, x_12); +return x_13; +} +else +{ +uint8_t x_14; +lean_dec(x_5); +lean_dec(x_4); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ return x_10; } else { -uint8_t x_11; +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; +} +} +} +else +{ +uint8_t x_18; lean_dec(x_5); lean_dec(x_4); -x_11 = !lean_is_exclusive(x_7); -if (x_11 == 0) +x_18 = !lean_is_exclusive(x_7); +if (x_18 == 0) { return x_7; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_7, 0); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_7, 0); +x_20 = lean_ctor_get(x_7, 1); +lean_inc(x_20); +lean_inc(x_19); lean_dec(x_7); -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; +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; } } } @@ -11277,7 +12906,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitFalseRec(lean_object { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_inc(x_1); -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1), 6, 1); +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_inferType___boxed), 6, 1); lean_closure_set(x_7, 0, x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___boxed), 6, 1); lean_closure_set(x_8, 0, x_7); @@ -11633,277 +13262,316 @@ return x_9; LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___lambda__2(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) { _start: { -lean_object* x_9; lean_object* x_51; uint8_t x_52; -x_51 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__13; -x_52 = l_Lean_Expr_isAppOf(x_2, x_51); -if (x_52 == 0) +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +x_9 = l_Lean_Compiler_ToLCNF_toLCNFType(x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_53; uint8_t x_54; -x_53 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__15; -x_54 = l_Lean_Expr_isAppOf(x_2, x_53); -lean_dec(x_2); -if (x_54 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_54; uint8_t x_55; +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_54 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__13; +x_55 = l_Lean_Expr_isAppOf(x_2, x_54); +if (x_55 == 0) { -lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_55 = lean_array_get_size(x_1); -x_56 = lean_unsigned_to_nat(5u); -x_57 = lean_nat_dec_lt(x_56, x_55); -lean_dec(x_55); +lean_object* x_56; uint8_t x_57; +x_56 = l_Lean_Compiler_ToLCNF_toLCNF_visitApp___closed__15; +x_57 = l_Lean_Expr_isAppOf(x_2, x_56); +lean_dec(x_2); if (x_57 == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; -x_59 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_58); -x_9 = x_59; -goto block_50; +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_array_get_size(x_1); +x_59 = lean_unsigned_to_nat(5u); +x_60 = lean_nat_dec_lt(x_59, x_58); +lean_dec(x_58); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; +x_62 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_61); +x_12 = x_62; +goto block_53; } else { -lean_object* x_60; -x_60 = lean_array_fget(x_1, x_56); -x_9 = x_60; -goto block_50; +lean_object* x_63; +x_63 = lean_array_fget(x_1, x_59); +x_12 = x_63; +goto block_53; } } else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_61 = lean_array_get_size(x_1); -x_62 = lean_unsigned_to_nat(3u); -x_63 = lean_nat_dec_lt(x_62, x_61); -lean_dec(x_61); -if (x_63 == 0) +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = lean_array_get_size(x_1); +x_65 = lean_unsigned_to_nat(3u); +x_66 = lean_nat_dec_lt(x_65, x_64); +lean_dec(x_64); +if (x_66 == 0) { -lean_object* x_64; lean_object* x_65; -x_64 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; -x_65 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_64); -x_9 = x_65; -goto block_50; +lean_object* x_67; lean_object* x_68; +x_67 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; +x_68 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_67); +x_12 = x_68; +goto block_53; } else { -lean_object* x_66; -x_66 = lean_array_fget(x_1, x_62); -x_9 = x_66; -goto block_50; +lean_object* x_69; +x_69 = lean_array_fget(x_1, x_65); +x_12 = x_69; +goto block_53; } } } else { -lean_object* x_67; lean_object* x_68; uint8_t x_69; +lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_dec(x_2); -x_67 = lean_array_get_size(x_1); -x_68 = lean_unsigned_to_nat(3u); -x_69 = lean_nat_dec_lt(x_68, x_67); -lean_dec(x_67); -if (x_69 == 0) +x_70 = lean_array_get_size(x_1); +x_71 = lean_unsigned_to_nat(3u); +x_72 = lean_nat_dec_lt(x_71, x_70); +lean_dec(x_70); +if (x_72 == 0) { -lean_object* x_70; lean_object* x_71; -x_70 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; -x_71 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_70); -x_9 = x_71; -goto block_50; +lean_object* x_73; lean_object* x_74; +x_73 = l_Lean_Compiler_ToLCNF_toLCNF_visitNoConfusion___lambda__2___closed__8; +x_74 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_73); +x_12 = x_74; +goto block_53; } else { -lean_object* x_72; -x_72 = lean_array_fget(x_1, x_68); -x_9 = x_72; -goto block_50; +lean_object* x_75; +x_75 = lean_array_fget(x_1, x_71); +x_12 = x_75; +goto block_53; } } -block_50: +block_53: { -lean_object* x_10; +lean_object* x_13; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_10 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_9, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_10) == 0) +x_13 = l_Lean_Compiler_ToLCNF_toLCNF_visit(x_12, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_13) == 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; -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_st_ref_get(x_7, x_12); -x_14 = lean_ctor_get(x_13, 1); +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_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_15 = lean_st_ref_get(x_5, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +x_16 = lean_st_ref_get(x_7, x_15); +x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); lean_dec(x_16); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_11); -x_19 = l_Lean_Compiler_InferType_inferType(x_11, x_18, x_6, x_7, x_17); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_19, 0); +x_18 = lean_st_ref_get(x_5, 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); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_3); -x_22 = l_Lean_Compiler_compatibleTypes(x_20, x_3); -if (x_22 == 0) -{ -uint8_t x_23; lean_object* x_24; -x_23 = 0; lean_inc(x_7); lean_inc(x_6); -x_24 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_11, x_23, x_5, x_6, x_7, x_21); -if (lean_obj_tag(x_24) == 0) +lean_inc(x_14); +x_22 = l_Lean_Compiler_InferType_inferType(x_14, x_21, x_6, x_7, x_20); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -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_object* x_23; lean_object* x_24; uint8_t x_25; +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); +lean_inc(x_10); +x_25 = l_Lean_Compiler_compatibleTypes(x_23, x_10); +if (x_25 == 0) +{ +uint8_t x_26; lean_object* x_27; +x_26 = 0; lean_inc(x_7); lean_inc(x_6); -x_27 = l_Lean_Compiler_mkLcCast___at_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___spec__1(x_25, x_3, x_4, x_5, x_6, x_7, x_26); +x_27 = l_Lean_Compiler_ToLCNF_mkAuxLetDecl(x_14, x_26, x_5, x_6, x_7, x_24); 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_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_unsigned_to_nat(6u); -x_31 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_28, x_1, x_30, x_4, x_5, x_6, x_7, x_29); -return x_31; +lean_inc(x_7); +lean_inc(x_6); +x_30 = l_Lean_Compiler_mkLcCast___at_Lean_Compiler_ToLCNF_toLCNF_visitEqRec___spec__1(x_28, x_10, x_4, x_5, x_6, x_7, x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +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_unsigned_to_nat(6u); +x_34 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_31, x_1, x_33, x_4, x_5, x_6, x_7, x_32); +return x_34; } else { -uint8_t x_32; +uint8_t x_35; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +x_35 = !lean_is_exclusive(x_30); +if (x_35 == 0) { -return x_27; +return x_30; } else { -lean_object* x_33; lean_object* x_34; lean_object* 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 = 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_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_30, 0); +x_37 = lean_ctor_get(x_30, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_30); +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; } } } else { -uint8_t x_36; +uint8_t x_39; +lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_24); -if (x_36 == 0) +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) { -return x_24; +return x_27; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_24, 0); -x_38 = lean_ctor_get(x_24, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_24); -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; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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_40; lean_object* x_41; -lean_dec(x_3); -x_40 = lean_unsigned_to_nat(6u); -x_41 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_11, x_1, x_40, x_4, x_5, x_6, x_7, x_21); -return x_41; +lean_object* x_43; lean_object* x_44; +lean_dec(x_10); +x_43 = lean_unsigned_to_nat(6u); +x_44 = l_Lean_Compiler_ToLCNF_toLCNF_mkOverApplication(x_14, x_1, x_43, x_4, x_5, x_6, x_7, x_24); +return x_44; } } else { -uint8_t x_42; -lean_dec(x_11); +uint8_t x_45; +lean_dec(x_14); +lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_42 = !lean_is_exclusive(x_19); -if (x_42 == 0) +x_45 = !lean_is_exclusive(x_22); +if (x_45 == 0) { -return x_19; +return x_22; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_19, 0); -x_44 = lean_ctor_get(x_19, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_19); -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; +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_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; } } } else { -uint8_t x_46; +uint8_t x_49; +lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_46 = !lean_is_exclusive(x_10); -if (x_46 == 0) +x_49 = !lean_is_exclusive(x_13); +if (x_49 == 0) { -return x_10; +return x_13; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_10, 0); -x_48 = lean_ctor_get(x_10, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_10); -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_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_13, 0); +x_51 = lean_ctor_get(x_13, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_13); +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_76; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_76 = !lean_is_exclusive(x_9); +if (x_76 == 0) +{ +return x_9; } +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_9, 0); +x_78 = lean_ctor_get(x_9, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_9); +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; } } } @@ -11928,7 +13596,7 @@ lean_inc(x_13); x_16 = l_Array_toSubarray___rarg(x_13, x_7, x_15); x_17 = l_Array_ofSubarray___rarg(x_16); x_18 = l_Lean_mkAppN(x_14, x_17); -x_19 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1), 6, 1); +x_19 = lean_alloc_closure((void*)(l_Lean_Meta_inferType___boxed), 6, 1); lean_closure_set(x_19, 0, x_18); x_20 = lean_alloc_closure((void*)(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___boxed), 6, 1); lean_closure_set(x_20, 0, x_19); @@ -11979,7 +13647,7 @@ static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__ 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_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__1; x_2 = l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(311u); +x_3 = lean_unsigned_to_nat(322u); x_4 = lean_unsigned_to_nat(42u); x_5 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12484,14 +14152,14 @@ lean_dec(x_4); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___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_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___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: { -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_3); -lean_dec(x_3); -x_9 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_1, x_2, x_8, x_4, x_5, x_6, x_7); -return x_9; +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_4); +lean_dec(x_4); +x_10 = l_Lean_Compiler_ToLCNF_toLCNF_visitCore___lambda__2(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -12905,13 +14573,13 @@ lean_dec(x_1); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF_visitCases___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_6); lean_dec(x_6); -x_12 = l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__2(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Compiler_ToLCNF_toLCNF_visitCases___lambda__1(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); return x_12; } } @@ -13085,18 +14753,32 @@ return x_8; static lean_object* _init_l_Lean_Compiler_ToLCNF_toLCNF___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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_ToLCNF_State_lctx___default___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_Compiler_ToLCNF_toLCNF___closed__2() { +_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_Compiler_ToLCNF_State_lctx___default___closed__7; x_2 = l_Lean_Compiler_ToLCNF_State_letFVars___default___closed__1; x_3 = lean_unsigned_to_nat(1u); -x_4 = l_Lean_Compiler_ToLCNF_State_cache___default___closed__1; -x_5 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_1); -lean_ctor_set(x_5, 2, x_2); -lean_ctor_set(x_5, 3, x_3); -lean_ctor_set(x_5, 4, x_4); -return x_5; +x_4 = l_Lean_Compiler_ToLCNF_toLCNF___closed__1; +x_5 = l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1; +x_6 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_2); +lean_ctor_set(x_6, 3, x_3); +lean_ctor_set(x_6, 4, x_4); +lean_ctor_set(x_6, 5, x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Compiler_ToLCNF_toLCNF(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -13107,7 +14789,7 @@ x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Compiler_ToLCNF_toLCNF___closed__1; +x_7 = l_Lean_Compiler_ToLCNF_toLCNF___closed__2; x_8 = lean_st_mk_ref(x_7, x_6); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); @@ -13262,8 +14944,22 @@ l_Lean_Compiler_ToLCNF_State_nextIdx___default = _init_l_Lean_Compiler_ToLCNF_St lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_nextIdx___default); l_Lean_Compiler_ToLCNF_State_cache___default___closed__1 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__1(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__1); +l_Lean_Compiler_ToLCNF_State_cache___default___closed__2 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__2(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__2); +l_Lean_Compiler_ToLCNF_State_cache___default___closed__3 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__3(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__3); +l_Lean_Compiler_ToLCNF_State_cache___default___closed__4 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__4(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__4); +l_Lean_Compiler_ToLCNF_State_cache___default___closed__5 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__5(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__5); +l_Lean_Compiler_ToLCNF_State_cache___default___closed__6 = _init_l_Lean_Compiler_ToLCNF_State_cache___default___closed__6(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default___closed__6); l_Lean_Compiler_ToLCNF_State_cache___default = _init_l_Lean_Compiler_ToLCNF_State_cache___default(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_cache___default); +l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1 = _init_l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_typeCache___default___closed__1); +l_Lean_Compiler_ToLCNF_State_typeCache___default = _init_l_Lean_Compiler_ToLCNF_State_typeCache___default(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_State_typeCache___default); l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1 = _init_l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__1); l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__2 = _init_l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__2(); @@ -13290,16 +14986,16 @@ l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__12 = _init_l_Lean_Compiler_ToL lean_mark_persistent(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__12); l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13 = _init_l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_liftMetaM___rarg___closed__13); +l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__1 = _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__1(); +l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2 = _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__2___closed__2(); +l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1 = _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1(); +lean_mark_persistent(l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNFType___spec__5___closed__1); l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__1 = _init_l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__1(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__1); l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__2 = _init_l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__2(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_mkAuxLetDecl___closed__2); l_Lean_Compiler_ToLCNF_mkLetUsingScope___closed__1 = _init_l_Lean_Compiler_ToLCNF_mkLetUsingScope___closed__1(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_mkLetUsingScope___closed__1); -l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__1 = _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__1(); -l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2 = _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__2(); -l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3 = _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3(); -lean_mark_persistent(l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__2___closed__3); l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__1 = _init_l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__1(); lean_mark_persistent(l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__1); l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__2 = _init_l_panic___at_Lean_Compiler_ToLCNF_toLCNF_visitCore___spec__5___closed__2(); @@ -13460,6 +15156,8 @@ l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__4 = _init_l_Le lean_mark_persistent(l_Lean_Compiler_ToLCNF_toLCNF_visitQuotLift___lambda__1___closed__4); l_Lean_Compiler_ToLCNF_toLCNF___closed__1 = _init_l_Lean_Compiler_ToLCNF_toLCNF___closed__1(); lean_mark_persistent(l_Lean_Compiler_ToLCNF_toLCNF___closed__1); +l_Lean_Compiler_ToLCNF_toLCNF___closed__2 = _init_l_Lean_Compiler_ToLCNF_toLCNF___closed__2(); +lean_mark_persistent(l_Lean_Compiler_ToLCNF_toLCNF___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/LCNFTypes.c b/stage0/stdlib/Lean/Compiler/LCNFTypes.c index 64d66448f298..6bc6e3b7bb2a 100644 --- a/stage0/stdlib/Lean/Compiler/LCNFTypes.c +++ b/stage0/stdlib/Lean/Compiler/LCNFTypes.c @@ -23,14 +23,12 @@ lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object* lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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_Lean_Compiler_instantiateLCNFTypeLevelParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___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 lean_is_marked_borrowed(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_getDeclLCNFType___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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_Compiler_toLCNFType___lambda__2___closed__2; LEAN_EXPORT uint8_t l_Lean_Expr_erased(lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__10; LEAN_EXPORT uint8_t l_Lean_Compiler_compatibleTypes(lean_object*, lean_object*); @@ -38,9 +36,7 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object*, lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__20; lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_toLCNFType___lambda__4___closed__2; size_t lean_usize_sub(size_t, size_t); -static lean_object* l_Lean_Compiler_toLCNFType___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_getDeclLCNFType___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -52,37 +48,41 @@ lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Std_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___lambda__2(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_Compiler_toLCNFType___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at_Lean_Compiler_getDeclLCNFType___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Compiler_getDeclLCNFType___spec__2(lean_object*, size_t, lean_object*); size_t lean_usize_shift_right(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg___boxed(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_Compiler_getDeclLCNFType___closed__12; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_getDeclLCNFType___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___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_Compiler_getDeclLCNFType___closed__13; static lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_getDeclLCNFType___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNFTypeExtState_types___default; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_getDeclLCNFType___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_isEqv___at_Lean_Compiler_compatibleTypes___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg(lean_object*, uint8_t, lean_object*, 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_Lean_Compiler_instantiateLCNFTypeLevelParams___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_Compiler_erasedExpr___closed__3; lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___lambda__4(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_Compiler_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_headBeta(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_instantiateLCNFTypeLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__7; static lean_object* l_Lean_Compiler_instInhabitedLCNFTypeExtState___closed__1; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1(lean_object*); uint8_t l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_insert___at_Lean_Core_instantiateTypeLevelParams___spec__1(lean_object*, lean_object*, lean_object*); @@ -97,6 +97,7 @@ LEAN_EXPORT uint8_t l_Lean_Compiler_isTypeFormerType(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Core_instantiateTypeLevelParams___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitApp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_getDeclLCNFType___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -104,6 +105,7 @@ static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__19; static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__5; static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__14; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_isClass_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_isEqv___at_Lean_Compiler_compatibleTypes___spec__1___boxed(lean_object*, lean_object*); @@ -127,7 +129,9 @@ static lean_object* l_Lean_Compiler_erasedExpr___closed__1; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__6; lean_object* l_Lean_Expr_eta(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___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*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_erasedExpr; uint8_t lean_usize_dec_le(size_t, size_t); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_LCNFTypes___hyg_49____closed__1; @@ -141,6 +145,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___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_Compiler_toLCNFType___lambda__1___closed__1; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* lean_expr_abstract(lean_object*, lean_object*); @@ -149,6 +154,7 @@ static lean_object* l_Lean_Compiler_anyTypeExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_getDeclLCNFType___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isTypeFormerType___boxed(lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__21; +lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__4; uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -167,9 +173,9 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_anyTypeExpr; LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitApp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNFTypeExtState_types___default___closed__3; +uint8_t lean_is_class(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_compatibleTypes___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_getDeclLCNFType___closed__9; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_markBorrowed(lean_object*); static lean_object* l_Lean_Compiler_anyTypeExpr___closed__2; @@ -368,390 +374,280 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(lean_object* x_1, size_t x_2, size_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_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg(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: { -uint8_t x_10; -x_10 = lean_usize_dec_lt(x_3, x_2); -if (x_10 == 0) +lean_object* x_10; +x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, 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_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -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_9); -return x_11; +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_20; -x_12 = lean_array_uget(x_1, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_12); -x_20 = l_Lean_Meta_isProp(x_12, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); lean_inc(x_12); -x_24 = l_Lean_Meta_isTypeFormer(x_12, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_24) == 0) +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +else { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_unbox(x_25); -lean_dec(x_25); -if (x_26 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_10); +if (x_15 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_12); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = l_Lean_Compiler_erasedExpr; -x_29 = l_Lean_Expr_app___override(x_4, x_28); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_13 = x_30; -x_14 = x_27; -goto block_19; +return x_10; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -lean_dec(x_24); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_32 = l_Lean_Compiler_toLCNFType(x_12, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_32) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_10); +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_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1(lean_object* x_1) { +_start: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Expr_app___override(x_4, x_33); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_13 = x_36; -x_14 = x_34; -goto block_19; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg___boxed), 9, 0); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___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) { +_start: { -uint8_t x_37; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) +lean_object* x_13; lean_object* x_14; +x_13 = lean_array_push(x_1, x_2); +x_14 = l_Lean_Compiler_toLCNFType_visitForall(x_3, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) { -return x_32; +uint8_t x_15; +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); +x_17 = l_Lean_Expr_forallE___override(x_4, x_6, x_16, x_5); +lean_ctor_set(x_14, 0, x_17); +return x_14; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_32, 0); -x_39 = lean_ctor_get(x_32, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_32); -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_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_dec(x_14); +x_20 = l_Lean_Expr_forallE___override(x_4, x_6, x_18, x_5); +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_41; -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_22; lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -x_41 = !lean_is_exclusive(x_24); -if (x_41 == 0) +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) { -return x_24; +return x_14; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_24, 0); -x_43 = lean_ctor_get(x_24, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_24); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +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_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; +} +} } } +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitForall___lambda__2(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) { +_start: +{ +uint8_t x_12; lean_object* x_13; +lean_inc(x_1); +x_12 = lean_is_marked_borrowed(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_13 = l_Lean_Compiler_toLCNFType(x_1, x_7, x_8, x_9, x_10, x_11); +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 = lean_expr_abstract(x_14, x_2); +lean_dec(x_14); +if (x_12 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Compiler_toLCNFType_visitForall___lambda__1(x_2, x_6, x_3, x_4, x_5, x_16, x_17, x_7, x_8, x_9, x_10, x_15); +return x_18; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_12); -x_45 = lean_ctor_get(x_20, 1); -lean_inc(x_45); -lean_dec(x_20); -x_46 = l_Lean_Compiler_erasedExpr; -x_47 = l_Lean_Expr_app___override(x_4, x_46); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_13 = x_48; -x_14 = x_45; -goto block_19; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = l_Lean_markBorrowed(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Compiler_toLCNFType_visitForall___lambda__1(x_2, x_6, x_3, x_4, x_5, x_19, x_20, x_7, x_8, x_9, x_10, x_15); +return x_21; } } else { -uint8_t x_49; -lean_dec(x_12); +uint8_t x_22; +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_49 = !lean_is_exclusive(x_20); -if (x_49 == 0) +lean_dec(x_3); +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) { -return x_20; +return x_13; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_20, 0); -x_51 = lean_ctor_get(x_20, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_20); -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_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +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; } } -block_19: -{ -lean_object* x_15; size_t x_16; size_t x_17; -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_3 = x_17; -x_4 = x_15; -x_9 = x_14; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitApp___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; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_get_size(x_1); -x_9 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_10 = 0; -x_11 = l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(x_1, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -return x_11; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; -} -} -else -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_11); -if (x_16 == 0) -{ -return x_11; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_11); -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_Lean_Compiler_toLCNFType_visitApp(lean_object* x_1, lean_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_Compiler_toLCNFType_visitForall(lean_object* x_1, 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: { -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_8; -x_8 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_2, x_1, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -case 4: +if (lean_obj_tag(x_1) == 7) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_1, 0); +lean_object* 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; +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, 1); +x_10 = lean_ctor_get(x_1, 2); lean_inc(x_10); +x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); lean_dec(x_1); -lean_inc(x_9); -x_11 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); +x_12 = lean_expr_instantiate_rev(x_9, x_2); +lean_dec(x_9); +x_13 = lean_box(x_11); +lean_inc(x_8); lean_inc(x_12); -if (lean_obj_tag(x_12) == 5) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Expr_const___override(x_9, x_10); -x_15 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_2, x_14, x_3, x_4, x_5, x_6, x_13); -lean_dec(x_2); +x_14 = lean_alloc_closure((void*)(l_Lean_Compiler_toLCNFType_visitForall___lambda__2___boxed), 11, 5); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_10); +lean_closure_set(x_14, 3, x_8); +lean_closure_set(x_14, 4, x_13); +x_15 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg(x_8, x_11, x_12, x_14, x_3, x_4, x_5, x_6, x_7); return x_15; } else { -uint8_t x_16; -lean_dec(x_12); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_16 = !lean_is_exclusive(x_11); -if (x_16 == 0) +lean_object* x_16; lean_object* x_17; +x_16 = lean_expr_instantiate_rev(x_1, x_2); +lean_dec(x_1); +x_17 = l_Lean_Compiler_toLCNFType(x_16, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_11, 0); -lean_dec(x_17); -x_18 = l_Lean_Compiler_anyTypeExpr; -lean_ctor_set(x_11, 0, x_18); -return x_11; +uint8_t x_18; +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, 0); +x_20 = lean_expr_abstract(x_19, x_2); +lean_dec(x_2); +lean_dec(x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_dec(x_11); -x_20 = l_Lean_Compiler_anyTypeExpr; -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; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +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_expr_abstract(x_21, x_2); +lean_dec(x_2); +lean_dec(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_22); +return x_24; } } else { -uint8_t x_22; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +uint8_t x_25; lean_dec(x_2); -x_22 = !lean_is_exclusive(x_11); -if (x_22 == 0) +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) { -return x_11; +return x_17; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_11, 0); -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_11); -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_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +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; } -default: -{ -lean_object* x_26; lean_object* x_27; -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_26 = l_Lean_Compiler_anyTypeExpr; -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_7); -return x_27; } } } @@ -785,65 +681,6 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg(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: -{ -lean_object* x_10; -x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) -{ -return x_10; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -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_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg___boxed), 9, 0); -return x_2; -} -} static lean_object* _init_l_Lean_Compiler_toLCNFType___lambda__1___closed__1() { _start: { @@ -1023,159 +860,17 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___lambda__2(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) { +static lean_object* _init_l_Lean_Compiler_toLCNFType___lambda__2___closed__1() { _start: { -lean_object* x_12; lean_object* x_13; -x_12 = lean_expr_instantiate1(x_1, x_2); -x_13 = l_Lean_Compiler_toLCNFType(x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(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; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_Compiler_toLCNFType___lambda__1___closed__1; -x_17 = lean_array_push(x_16, x_2); -x_18 = lean_expr_abstract(x_15, x_17); -lean_dec(x_17); -lean_dec(x_15); -x_19 = l_Lean_Expr_forallE___override(x_3, x_5, x_18, x_4); -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; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = l_Lean_Compiler_toLCNFType___lambda__1___closed__1; -x_23 = lean_array_push(x_22, x_2); -x_24 = lean_expr_abstract(x_20, x_23); -lean_dec(x_23); -lean_dec(x_20); -x_25 = l_Lean_Expr_forallE___override(x_3, x_5, x_24, x_4); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_21); -return x_26; +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; } } -else -{ -uint8_t x_27; -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_27 = !lean_is_exclusive(x_13); -if (x_27 == 0) -{ -return x_13; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -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 = 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_Compiler_toLCNFType___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) { -_start: -{ -uint8_t x_11; lean_object* x_12; -lean_inc(x_1); -x_11 = lean_is_marked_borrowed(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = l_Lean_Compiler_toLCNFType(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -if (x_11 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -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_box(0); -x_16 = l_Lean_Compiler_toLCNFType___lambda__2(x_2, x_5, x_3, x_4, x_13, x_15, x_6, x_7, x_8, x_9, x_14); -lean_dec(x_2); -return x_16; -} -else -{ -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_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_Lean_markBorrowed(x_17); -x_20 = lean_box(0); -x_21 = l_Lean_Compiler_toLCNFType___lambda__2(x_2, x_5, x_3, x_4, x_19, x_20, x_6, x_7, x_8, x_9, x_18); -lean_dec(x_2); -return x_21; -} -} -else -{ -uint8_t 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_3); -lean_dec(x_2); -x_22 = !lean_is_exclusive(x_12); -if (x_22 == 0) -{ -return x_12; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_12, 0); -x_24 = lean_ctor_get(x_12, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_12); -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_Compiler_toLCNFType___lambda__4___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_Lean_Compiler_toLCNFType___lambda__4___closed__2() { -_start: +static lean_object* _init_l_Lean_Compiler_toLCNFType___lambda__2___closed__2() { +_start: { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_levelZero; @@ -1183,7 +878,7 @@ x_2 = l_Lean_Expr_sort___override(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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; @@ -1205,7 +900,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l_Lean_Compiler_toLCNFType___lambda__4___closed__1; +x_11 = l_Lean_Compiler_toLCNFType___lambda__2___closed__1; x_12 = l_Lean_Compiler_toLCNFType_visitApp(x_9, x_11, x_3, x_4, x_5, x_6, x_10); return x_12; } @@ -1251,7 +946,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_8, 1); lean_inc(x_21); lean_dec(x_8); -x_22 = l_Lean_Compiler_toLCNFType___lambda__4___closed__1; +x_22 = l_Lean_Compiler_toLCNFType___lambda__2___closed__1; x_23 = l_Lean_Compiler_toLCNFType_visitApp(x_9, x_22, x_3, x_4, x_5, x_6, x_21); return x_23; } @@ -1263,7 +958,7 @@ lean_inc(x_24); lean_dec(x_8); x_25 = lean_unsigned_to_nat(0u); x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_9, x_25); -x_27 = l_Lean_Compiler_toLCNFType___lambda__4___closed__2; +x_27 = l_Lean_Compiler_toLCNFType___lambda__2___closed__2; lean_inc(x_26); x_28 = lean_mk_array(x_26, x_27); x_29 = lean_unsigned_to_nat(1u); @@ -1294,213 +989,585 @@ lean_closure_set(x_38, 0, x_34); lean_closure_set(x_38, 1, x_35); lean_closure_set(x_38, 2, x_33); lean_closure_set(x_38, 3, x_37); -x_39 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg(x_33, x_36, x_34, x_38, x_3, x_4, x_5, x_6, x_32); +x_39 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg(x_33, x_36, x_34, x_38, x_3, x_4, x_5, x_6, x_32); return x_39; } case 7: { -lean_object* x_40; 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_40; lean_object* x_41; lean_object* x_42; x_40 = lean_ctor_get(x_8, 1); lean_inc(x_40); lean_dec(x_8); -x_41 = lean_ctor_get(x_9, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_9, 1); -lean_inc(x_42); -x_43 = lean_ctor_get(x_9, 2); -lean_inc(x_43); -x_44 = lean_ctor_get_uint8(x_9, sizeof(void*)*3 + 8); -lean_dec(x_9); -x_45 = lean_box(x_44); -lean_inc(x_41); -lean_inc(x_42); -x_46 = lean_alloc_closure((void*)(l_Lean_Compiler_toLCNFType___lambda__3___boxed), 10, 4); -lean_closure_set(x_46, 0, x_42); -lean_closure_set(x_46, 1, x_43); -lean_closure_set(x_46, 2, x_41); -lean_closure_set(x_46, 3, x_45); -x_47 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg(x_41, x_44, x_42, x_46, x_3, x_4, x_5, x_6, x_40); -return x_47; +x_41 = l_Lean_Compiler_toLCNFType___lambda__2___closed__1; +x_42 = l_Lean_Compiler_toLCNFType_visitForall(x_9, x_41, x_3, x_4, x_5, x_6, x_40); +return x_42; } default: { -uint8_t x_48; +uint8_t x_43; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_48 = !lean_is_exclusive(x_8); -if (x_48 == 0) +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_8, 0); +lean_dec(x_44); +x_45 = l_Lean_Compiler_anyTypeExpr; +lean_ctor_set(x_8, 0, x_45); +return x_8; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_8, 1); +lean_inc(x_46); +lean_dec(x_8); +x_47 = l_Lean_Compiler_anyTypeExpr; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +else +{ +uint8_t x_49; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_49 = !lean_is_exclusive(x_8); +if (x_49 == 0) +{ +return x_8; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_8, 0); +x_51 = lean_ctor_get(x_8, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_8); +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_Lean_Compiler_toLCNFType(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_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_7 = l_Lean_Meta_isProp(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_box(0); +x_12 = l_Lean_Compiler_toLCNFType___lambda__2(x_1, x_11, x_2, x_3, x_4, x_5, x_10); +return x_12; +} +else +{ +uint8_t x_13; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_7); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_7, 0); +lean_dec(x_14); +x_15 = l_Lean_Compiler_erasedExpr; +lean_ctor_set(x_7, 0, x_15); +return x_7; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 1); +lean_inc(x_16); +lean_dec(x_7); +x_17 = l_Lean_Compiler_erasedExpr; +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 +{ +uint8_t x_19; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(lean_object* x_1, size_t x_2, size_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: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +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_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_20; +x_12 = lean_array_uget(x_1, x_3); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_12); +x_20 = l_Lean_Meta_isProp(x_12, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_8, 0); -lean_dec(x_49); -x_50 = l_Lean_Compiler_anyTypeExpr; -lean_ctor_set(x_8, 0, x_50); +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_12); +x_24 = l_Lean_Meta_isTypeFormer(x_12, x_5, x_6, x_7, x_8, x_23); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_12); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = l_Lean_Compiler_erasedExpr; +x_29 = l_Lean_Expr_app___override(x_4, x_28); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_13 = x_30; +x_14 = x_27; +goto block_19; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_24, 1); +lean_inc(x_31); +lean_dec(x_24); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_32 = l_Lean_Compiler_toLCNFType(x_12, x_5, x_6, x_7, x_8, x_31); +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, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Expr_app___override(x_4, x_33); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_13 = x_36; +x_14 = x_34; +goto block_19; +} +else +{ +uint8_t x_37; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_37 = !lean_is_exclusive(x_32); +if (x_37 == 0) +{ +return x_32; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_32, 0); +x_39 = lean_ctor_get(x_32, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_32); +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 +{ +uint8_t x_41; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_41 = !lean_is_exclusive(x_24); +if (x_41 == 0) +{ +return x_24; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_24, 0); +x_43 = lean_ctor_get(x_24, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_24); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +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; +lean_dec(x_12); +x_45 = lean_ctor_get(x_20, 1); +lean_inc(x_45); +lean_dec(x_20); +x_46 = l_Lean_Compiler_erasedExpr; +x_47 = l_Lean_Expr_app___override(x_4, x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_13 = x_48; +x_14 = x_45; +goto block_19; +} +} +else +{ +uint8_t x_49; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_49 = !lean_is_exclusive(x_20); +if (x_49 == 0) +{ +return x_20; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_20, 0); +x_51 = lean_ctor_get(x_20, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_20); +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_19: +{ +lean_object* x_15; size_t x_16; size_t x_17; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = 1; +x_17 = lean_usize_add(x_3, x_16); +x_3 = x_17; +x_4 = x_15; +x_9 = x_14; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitApp___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; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_get_size(x_1); +x_9 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_10 = 0; +x_11 = l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(x_1, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) +{ +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +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_Lean_Compiler_toLCNFType_visitApp(lean_object* x_1, 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: +{ +switch (lean_obj_tag(x_1)) { +case 1: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_2, x_1, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); return x_8; } -else +case 4: { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_8, 1); -lean_inc(x_51); -lean_dec(x_8); -x_52 = l_Lean_Compiler_anyTypeExpr; -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; -} -} -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_dec(x_1); +lean_inc(x_9); +x_11 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); +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) == 5) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Expr_const___override(x_9, x_10); +x_15 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_2, x_14, x_3, x_4, x_5, x_6, x_13); +lean_dec(x_2); +return x_15; } else { -uint8_t x_54; +uint8_t x_16; +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_54 = !lean_is_exclusive(x_8); -if (x_54 == 0) +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -return x_8; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_11, 0); +lean_dec(x_17); +x_18 = l_Lean_Compiler_anyTypeExpr; +lean_ctor_set(x_11, 0, x_18); +return x_11; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_8, 0); -x_56 = lean_ctor_get(x_8, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_8); -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; -} -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = l_Lean_Compiler_anyTypeExpr; +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; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType(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_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_7 = l_Lean_Meta_isProp(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = lean_box(0); -x_12 = l_Lean_Compiler_toLCNFType___lambda__4(x_1, x_11, x_2, x_3, x_4, x_5, x_10); -return x_12; } else { -uint8_t x_13; +uint8_t x_22; +lean_dec(x_10); +lean_dec(x_9); +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_13 = !lean_is_exclusive(x_7); -if (x_13 == 0) +x_22 = !lean_is_exclusive(x_11); +if (x_22 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_7, 0); -lean_dec(x_14); -x_15 = l_Lean_Compiler_erasedExpr; -lean_ctor_set(x_7, 0, x_15); -return x_7; +return x_11; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_7, 1); -lean_inc(x_16); -lean_dec(x_7); -x_17 = l_Lean_Compiler_erasedExpr; -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; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_11, 0); +x_24 = lean_ctor_get(x_11, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_11); +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 +default: { -uint8_t x_19; +lean_object* x_26; lean_object* x_27; +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_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; +x_26 = l_Lean_Compiler_anyTypeExpr; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_7); +return x_27; } } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___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_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___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_2); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); lean_dec(x_2); -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_1); -return x_12; +x_11 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType_visitForall___spec__1___rarg(x_1, 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_Compiler_toLCNFType_visitApp___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_Compiler_toLCNFType_visitForall___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: { -lean_object* x_8; -x_8 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_1); -return x_8; +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = l_Lean_Compiler_toLCNFType_visitForall___lambda__1(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); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___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_Compiler_toLCNFType_visitForall___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: { -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = l_Lean_Meta_withLocalDecl___at_Lean_Compiler_toLCNFType___spec__2___rarg(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Compiler_toLCNFType_visitForall___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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) { @@ -1514,26 +1581,26 @@ lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___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) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = l_Lean_Compiler_toLCNFType___lambda__2(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_6); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_toLCNFType_visitApp___spec__1(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_1); -return x_13; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_toLCNFType___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_EXPORT lean_object* l_Lean_Compiler_toLCNFType_visitApp___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: { -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_4); -lean_dec(x_4); -x_12 = l_Lean_Compiler_toLCNFType___lambda__3(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; +lean_object* x_8; +x_8 = l_Lean_Compiler_toLCNFType_visitApp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_8; } } LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_getDeclLCNFType___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -2362,7 +2429,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = lean_box(0); x_2 = l_Lean_Compiler_getDeclLCNFType___closed__2; x_3 = l_Lean_Compiler_getDeclLCNFType___closed__7; -x_4 = l_Lean_Compiler_toLCNFType___lambda__4___closed__1; +x_4 = l_Lean_Compiler_toLCNFType___lambda__2___closed__1; x_5 = lean_unsigned_to_nat(0u); x_6 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_6, 0, x_2); @@ -3837,6 +3904,103 @@ x_3 = lean_box(x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_isClass_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_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_5) == 4) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_st_ref_get(x_3, x_4); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +lean_inc(x_6); +x_11 = lean_is_class(x_10, x_6); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_6); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else +{ +lean_object* x_13; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_7, 0, x_13); +return x_7; +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_7); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_6); +x_17 = lean_is_class(x_16, x_6); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_6); +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_15); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_6); +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; +} +} +} +else +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_5); +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_4); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_isClass_x3f___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_Compiler_isClass_x3f(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} lean_object* initialize_Init(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_BorrowedAnnotation(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_InferType(uint8_t builtin, lean_object*); @@ -3893,10 +4057,10 @@ l_Lean_Compiler_anyTypeExpr = _init_l_Lean_Compiler_anyTypeExpr(); lean_mark_persistent(l_Lean_Compiler_anyTypeExpr); l_Lean_Compiler_toLCNFType___lambda__1___closed__1 = _init_l_Lean_Compiler_toLCNFType___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Compiler_toLCNFType___lambda__1___closed__1); -l_Lean_Compiler_toLCNFType___lambda__4___closed__1 = _init_l_Lean_Compiler_toLCNFType___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Compiler_toLCNFType___lambda__4___closed__1); -l_Lean_Compiler_toLCNFType___lambda__4___closed__2 = _init_l_Lean_Compiler_toLCNFType___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Compiler_toLCNFType___lambda__4___closed__2); +l_Lean_Compiler_toLCNFType___lambda__2___closed__1 = _init_l_Lean_Compiler_toLCNFType___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Compiler_toLCNFType___lambda__2___closed__1); +l_Lean_Compiler_toLCNFType___lambda__2___closed__2 = _init_l_Lean_Compiler_toLCNFType___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Compiler_toLCNFType___lambda__2___closed__2); l_Std_PersistentHashMap_findAux___at_Lean_Compiler_getDeclLCNFType___spec__2___closed__1 = _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_getDeclLCNFType___spec__2___closed__1(); l_Std_PersistentHashMap_findAux___at_Lean_Compiler_getDeclLCNFType___spec__2___closed__2 = _init_l_Std_PersistentHashMap_findAux___at_Lean_Compiler_getDeclLCNFType___spec__2___closed__2(); l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_getDeclLCNFType___spec__5___closed__1 = _init_l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_getDeclLCNFType___spec__5___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/Main.c b/stage0/stdlib/Lean/Compiler/Main.c index 835f1881e3ed..0d31c830fc66 100644 --- a/stage0/stdlib/Lean/Compiler/Main.c +++ b/stage0/stdlib/Lean/Compiler/Main.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.Main -// Imports: Init Lean.Compiler.Decl Lean.Compiler.TerminalCases Lean.Compiler.CSE Lean.Compiler.Stage1 Lean.Compiler.Simp +// Imports: Init Lean.Compiler.Decl Lean.Compiler.TerminalCases Lean.Compiler.CSE Lean.Compiler.Stage1 Lean.Compiler.Simp Lean.Compiler.PullLocalDecls #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* l_Lean_Compiler_toDecl(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__1; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); @@ -32,11 +33,11 @@ uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__15; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__22; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__17; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__14; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__5; +lean_object* l_Lean_Compiler_Decl_pullInstances___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__14; LEAN_EXPORT lean_object* l_Lean_Compiler_checkpoint___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -46,6 +47,8 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spe static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__3; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__15; lean_object* l_Lean_Compiler_Decl_mapValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -54,28 +57,34 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2___closed__1; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__5; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Compiler_compile___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__12; -LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at_Lean_Compiler_shouldGenerateCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1; static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_compile(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__6; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at_Lean_Compiler_shouldGenerateCode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__13; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__21; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__23; LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__13; -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3; +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4; +lean_object* l_Nat_repr(lean_object*); +static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -83,18 +92,18 @@ uint8_t lean_is_matcher(lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Compiler_Decl_cse___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__4; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_checkpoint___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_compile_stage1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1; static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode_isCompIrrelevant(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_checkpoint(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_checkpoint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__9; lean_object* l_Lean_Meta_isTypeFormerType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -106,25 +115,36 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1(lean_obje LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__16; lean_object* l_Lean_Compiler_Decl_simp(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_Decl_check(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_Decl_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__11; static lean_object* l_Lean_Compiler_compile___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_Decl_pullLocalDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1; lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); uint8_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(lean_object*, uint8_t, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__19; +static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__11; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2; lean_object* l_Lean_Compiler_saveStage1Decls(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4; lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_shouldGenerateCode(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_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_checkpoint___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__12; static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_checkpoint___spec__1___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -136,20 +156,32 @@ static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__3; extern lean_object* l_Lean_Expr_instBEqExpr; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Compiler_compile___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Compiler_compileStage1Impl___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__8(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_checkpoint___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_compile___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__2; +lean_object* l_Lean_Compiler_getLCNFSize(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__10; static lean_object* l_Lean_Compiler_shouldGenerateCode___closed__20; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1; static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__8; +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__4; +static lean_object* l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5___boxed(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_Lean_Compiler_shouldGenerateCode_isCompIrrelevant(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: @@ -1064,7 +1096,7 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___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; @@ -1205,7 +1237,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3(lean_object* x_1, uint8_t 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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___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, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -1215,6 +1247,7 @@ if (x_10 == 0) lean_object* x_11; lean_dec(x_8); lean_dec(x_7); +lean_dec(x_2); lean_dec(x_1); x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_6); @@ -1280,6 +1313,7 @@ x_40 = lean_ctor_get(x_37, 1); lean_inc(x_40); lean_dec(x_37); lean_inc(x_8); +lean_inc(x_2); x_41 = l_Lean_Compiler_Decl_check(x_12, x_2, x_36, x_8, x_40); if (lean_obj_tag(x_41) == 0) { @@ -1297,6 +1331,7 @@ else uint8_t x_44; lean_dec(x_8); lean_dec(x_7); +lean_dec(x_2); lean_dec(x_1); x_44 = !lean_is_exclusive(x_41); if (x_44 == 0) @@ -1362,6 +1397,7 @@ x_65 = lean_ctor_get(x_64, 1); lean_inc(x_65); lean_dec(x_64); lean_inc(x_8); +lean_inc(x_2); x_66 = l_Lean_Compiler_Decl_check(x_12, x_2, x_36, x_8, x_65); if (lean_obj_tag(x_66) == 0) { @@ -1379,6 +1415,7 @@ else uint8_t x_69; lean_dec(x_8); lean_dec(x_7); +lean_dec(x_2); lean_dec(x_1); x_69 = !lean_is_exclusive(x_66); if (x_69 == 0) @@ -1416,7 +1453,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_checkpoint(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_Compiler_checkpoint(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; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; @@ -1497,38 +1534,32 @@ return x_6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___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) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_2); -lean_dec(x_2); -x_8 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___lambda__1(x_1, x_7, x_3, x_4, x_5, x_6); +lean_object* x_7; +x_7 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); -return x_8; +return x_7; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___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) { _start: { -uint8_t x_10; size_t x_11; size_t x_12; lean_object* x_13; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = lean_unbox_usize(x_4); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_4); lean_dec(x_4); -x_12 = lean_unbox_usize(x_5); +x_11 = lean_unbox_usize(x_5); lean_dec(x_5); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3(x_1, x_10, x_3, x_11, x_12, x_6, x_7, x_8, x_9); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3(x_1, x_2, x_3, x_10, x_11, x_6, x_7, x_8, x_9); lean_dec(x_3); -return x_13; +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Compiler_checkpoint___boxed(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_Compiler_checkpoint(x_1, x_2, x_7, x_4, x_5, x_6); +lean_object* x_7; +x_7 = l_Lean_Compiler_checkpoint(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_2); -return x_8; +return x_7; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -1676,83 +1707,226 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__3(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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_2, x_1); -if (x_7 == 0) -{ -lean_object* x_8; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_st_ref_get(x_3, x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); lean_dec(x_5); -lean_dec(x_4); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_3); -lean_ctor_set(x_8, 1, x_6); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_checkpoint___spec__1___closed__1; +x_8 = lean_st_ref_get(x_7, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_2, 2); +x_12 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_10, x_11, x_1); +lean_dec(x_10); +x_13 = lean_box(x_12); +lean_ctor_set(x_8, 0, x_13); return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_array_uget(x_3, x_2); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_3, x_2, x_10); -lean_inc(x_5); -lean_inc(x_4); -x_12 = l_Lean_Compiler_Decl_simp(x_9, x_4, x_5, x_6); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_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_12); -x_15 = 1; -x_16 = lean_usize_add(x_2, x_15); -x_17 = lean_array_uset(x_11, x_2, x_13); -x_2 = x_16; -x_3 = x_17; -x_6 = x_14; +lean_dec(x_8); +x_16 = lean_ctor_get(x_2, 2); +x_17 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_14, x_16, x_1); +lean_dec(x_14); +x_18 = lean_box(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_15); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Compiler_compileStage1Impl___spec__4(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 +{ +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; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} goto _start; } else { -uint8_t x_19; -lean_dec(x_11); -lean_dec(x_5); -lean_dec(x_4); -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +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_alloc_ctor(4, 1, 0); +lean_ctor_set(x_11, 0, x_9); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_12; +lean_object* x_6; 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, 5); +x_7 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_3, x_4, x_5); +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_st_ref_take(x_4, x_9); +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_is_exclusive(x_11); +if (x_13 == 0) +{ +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; uint8_t x_21; +x_14 = lean_ctor_get(x_11, 3); +x_15 = l_Lean_Compiler_shouldGenerateCode___closed__9; +x_16 = 0; +x_17 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_8); +lean_ctor_set(x_17, 2, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*3, x_16); +lean_inc(x_6); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_6); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Std_PersistentArray_push___rarg(x_14, x_18); +lean_ctor_set(x_11, 3, x_19); +x_20 = lean_st_ref_set(x_4, x_11, x_12); +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 = lean_box(0); +lean_ctor_set(x_20, 0, x_23); +return x_20; } 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; +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 = 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; 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; 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_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +x_29 = lean_ctor_get(x_11, 2); +x_30 = lean_ctor_get(x_11, 3); +x_31 = lean_ctor_get(x_11, 4); +x_32 = lean_ctor_get(x_11, 5); +x_33 = lean_ctor_get(x_11, 6); +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_dec(x_11); +x_34 = l_Lean_Compiler_shouldGenerateCode___closed__9; +x_35 = 0; +x_36 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_8); +lean_ctor_set(x_36, 2, x_34); +lean_ctor_set_uint8(x_36, sizeof(void*)*3, x_35); +lean_inc(x_6); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_6); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Std_PersistentArray_push___rarg(x_30, x_37); +x_39 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_39, 0, x_27); +lean_ctor_set(x_39, 1, x_28); +lean_ctor_set(x_39, 2, x_29); +lean_ctor_set(x_39, 3, x_38); +lean_ctor_set(x_39, 4, x_31); +lean_ctor_set(x_39, 5, x_32); +lean_ctor_set(x_39, 6, x_33); +x_40 = lean_st_ref_set(x_4, x_39, x_12); +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; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Decl_cse___lambda__1), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Decl_pullInstances___lambda__1___boxed), 6, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4(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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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) { _start: { uint8_t x_7; @@ -1773,10 +1947,10 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ x_9 = lean_array_uget(x_3, x_2); x_10 = lean_unsigned_to_nat(0u); x_11 = lean_array_uset(x_3, x_2, x_10); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1; +x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1; lean_inc(x_5); lean_inc(x_4); -x_13 = l_Lean_Compiler_Decl_mapValue(x_9, x_12, x_4, x_5, x_6); +x_13 = l_Lean_Compiler_Decl_pullLocalDecls(x_9, x_12, x_4, x_5, x_6); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; @@ -1821,33 +1995,523 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_eq(x_2, x_3); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_array_uget(x_1, x_2); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_9); -x_10 = l_Lean_Compiler_shouldGenerateCode(x_9, x_5, x_6, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Decl_cse___lambda__1), 5, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_2, x_1); +if (x_7 == 0) { -lean_object* x_13; size_t x_14; size_t x_15; -lean_dec(x_9); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); +lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_3); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_uget(x_3, x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_array_uset(x_3, x_2, x_10); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1; +lean_inc(x_5); +lean_inc(x_4); +x_13 = l_Lean_Compiler_Decl_mapValue(x_9, x_12, x_4, x_5, x_6); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; size_t 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); +lean_dec(x_13); +x_16 = 1; +x_17 = lean_usize_add(x_2, x_16); +x_18 = lean_array_uset(x_11, x_2, x_14); +x_2 = x_17; +x_3 = x_18; +x_6 = x_15; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +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); +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_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__8(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_2, x_1); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_3); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_array_uget(x_3, x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_array_uset(x_3, x_2, x_10); +lean_inc(x_5); +lean_inc(x_4); +x_12 = l_Lean_Compiler_Decl_simp(x_9, x_4, x_5, x_6); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +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 = 1; +x_16 = lean_usize_add(x_2, x_15); +x_17 = lean_array_uset(x_11, x_2, x_13); +x_2 = x_16; +x_3 = x_17; +x_6 = x_14; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("stat", 4); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(": ", 2); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_2, 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; +lean_dec(x_4); +x_9 = lean_array_uget(x_1, x_2); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2; +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3(x_10, x_5, x_6, x_7); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +lean_dec(x_9); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = 1; +x_16 = lean_usize_add(x_2, x_15); +x_17 = lean_box(0); +x_2 = x_16; +x_4 = x_17; +x_7 = x_14; +goto _start; +} +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_9, 0); +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_5); +x_22 = l_Lean_Compiler_getLCNFSize(x_21, x_5, x_6, x_19); +if (lean_obj_tag(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; size_t x_38; size_t x_39; +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_alloc_ctor(4, 1, 0); +lean_ctor_set(x_25, 0, x_20); +x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__11; +x_27 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4; +x_29 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Nat_repr(x_23); +x_31 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_29); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_26); +x_35 = l_Lean_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5(x_10, x_34, x_5, x_6, x_24); +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 = 1; +x_39 = lean_usize_add(x_2, x_38); +x_2 = x_39; +x_4 = x_36; +x_7 = x_37; +goto _start; +} +else +{ +uint8_t x_41; +lean_dec(x_20); +lean_dec(x_6); +lean_dec(x_5); +x_41 = !lean_is_exclusive(x_22); +if (x_41 == 0) +{ +return x_22; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_22, 0); +x_43 = lean_ctor_get(x_22, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_22); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +else +{ +lean_object* x_45; +lean_dec(x_6); +lean_dec(x_5); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_4); +lean_ctor_set(x_45, 1, x_7); +return x_45; +} +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("jp", 2); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Compiler_shouldGenerateCode___closed__8; +x_2 = l_Lean_Compiler_shouldGenerateCode___closed__9; +x_3 = lean_unsigned_to_nat(1u); +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_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_2, x_3); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_4); +x_9 = lean_array_uget(x_1, x_2); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2; +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3(x_16, x_5, x_6, x_7); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_9); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_10 = x_21; +x_11 = x_20; +goto block_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_29; lean_object* x_30; lean_object* x_31; +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_dec(x_17); +x_23 = lean_ctor_get(x_9, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_9, 3); +lean_inc(x_24); +lean_dec(x_9); +x_25 = lean_st_ref_get(x_6, x_22); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3; +x_28 = lean_st_mk_ref(x_27, 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_6); +lean_inc(x_5); +lean_inc(x_29); +x_31 = l_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints(x_24, x_29, x_5, x_6, x_30); +if (lean_obj_tag(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_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_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_st_ref_get(x_6, x_33); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_st_ref_get(x_29, x_35); +lean_dec(x_29); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_38, 0, x_23); +x_39 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__11; +x_40 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4; +x_42 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_array_to_list(lean_box(0), x_32); +x_44 = lean_box(0); +x_45 = l_List_mapTRAux___at_Lean_Compiler_compileStage1Impl___spec__4(x_43, x_44); +x_46 = l_Lean_MessageData_ofList(x_45); +lean_dec(x_45); +x_47 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_47, 0, x_42); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_39); +x_49 = l_Lean_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5(x_16, x_48, x_5, x_6, x_37); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_10 = x_50; +x_11 = x_51; +goto block_15; +} +else +{ +uint8_t x_52; +lean_dec(x_29); +lean_dec(x_23); +lean_dec(x_6); +lean_dec(x_5); +x_52 = !lean_is_exclusive(x_31); +if (x_52 == 0) +{ +return x_31; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_31, 0); +x_54 = lean_ctor_get(x_31, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_31); +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; +} +} +} +block_15: +{ +size_t x_12; size_t x_13; +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_2 = x_13; +x_4 = x_10; +x_7 = x_11; +goto _start; +} +} +else +{ +lean_object* x_56; +lean_dec(x_6); +lean_dec(x_5); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_4); +lean_ctor_set(x_56, 1, x_7); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__11(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_2, x_3); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_array_uget(x_1, x_2); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_9); +x_10 = l_Lean_Compiler_shouldGenerateCode(x_9, x_5, x_6, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; size_t x_14; size_t x_15; +lean_dec(x_9); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); x_14 = 1; x_15 = lean_usize_add(x_2, x_14); x_2 = x_15; @@ -1929,40 +2593,63 @@ return x_3; static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__3() { _start: { +uint8_t x_1; uint8_t x_2; lean_object* x_3; +x_1 = 1; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 0, 2); +lean_ctor_set_uint8(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__4() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("terminalCases", 13); return x_1; } } -static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__3; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__6() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(0, 0, 2); +lean_ctor_set_uint8(x_2, 0, x_1); +lean_ctor_set_uint8(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("simp", 4); +x_1 = lean_mk_string_from_bytes("pullInstances", 13); return x_1; } } -static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9() { _start: { lean_object* x_1; @@ -1970,12 +2657,30 @@ x_1 = lean_mk_string_from_bytes("cse", 3); return x_1; } } -static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("simp", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -1993,14 +2698,14 @@ lean_inc(x_3); x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__1(x_7, x_8, x_1, x_3, x_4, x_5); if (lean_obj_tag(x_9) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +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_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_Compiler_compileStage1Impl___lambda__1___closed__2; -x_13 = 0; +x_13 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__3; lean_inc(x_4); lean_inc(x_3); x_14 = l_Lean_Compiler_checkpoint(x_12, x_10, x_13, x_3, x_4, x_11); @@ -2018,297 +2723,593 @@ lean_inc(x_3); x_18 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2(x_17, x_8, x_10, x_3, x_4, x_15); if (lean_obj_tag(x_18) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +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_Compiler_compileStage1Impl___lambda__1___closed__4; -x_22 = 1; +x_21 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5; +x_22 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__6; lean_inc(x_4); lean_inc(x_3); x_23 = l_Lean_Compiler_checkpoint(x_21, x_19, x_22, x_3, x_4, x_20); if (lean_obj_tag(x_23) == 0) { -lean_object* x_24; lean_object* x_25; size_t x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_array_get_size(x_19); -x_26 = lean_usize_of_nat(x_25); -lean_dec(x_25); -lean_inc(x_4); -lean_inc(x_3); -x_27 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__3(x_26, x_8, x_19, x_3, x_4, x_24); -if (lean_obj_tag(x_27) == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_array_get_size(x_19); +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_nat_dec_lt(x_26, x_25); +if (x_27 == 0) +{ +x_28 = x_24; +goto block_110; +} +else +{ +uint8_t x_111; +x_111 = lean_nat_dec_le(x_25, x_25); +if (x_111 == 0) +{ +x_28 = x_24; +goto block_110; +} +else +{ +size_t x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_usize_of_nat(x_25); +x_113 = lean_box(0); +lean_inc(x_4); +lean_inc(x_3); +x_114 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10(x_19, x_8, x_112, x_113, x_3, x_4, x_24); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_28 = x_115; +goto block_110; +} +else +{ +uint8_t x_116; +lean_dec(x_25); +lean_dec(x_19); +lean_dec(x_4); +lean_dec(x_3); +x_116 = !lean_is_exclusive(x_114); +if (x_116 == 0) +{ +return x_114; +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_114, 0); +x_118 = lean_ctor_get(x_114, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_114); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +} +} +} +block_110: +{ +size_t x_29; lean_object* x_30; +x_29 = lean_usize_of_nat(x_25); +lean_dec(x_25); +lean_inc(x_4); +lean_inc(x_3); +x_30 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6(x_29, x_8, x_19, x_3, x_4, x_28); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +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 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8; +lean_inc(x_4); +lean_inc(x_3); +x_34 = l_Lean_Compiler_checkpoint(x_33, x_31, x_22, x_3, x_4, x_32); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; size_t x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_array_get_size(x_31); +x_37 = lean_usize_of_nat(x_36); +lean_dec(x_36); +lean_inc(x_4); +lean_inc(x_3); +x_38 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7(x_37, x_8, x_31, x_3, x_4, x_35); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +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 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10; +lean_inc(x_4); +lean_inc(x_3); +x_42 = l_Lean_Compiler_checkpoint(x_41, x_39, x_22, x_3, x_4, x_40); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; size_t x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_array_get_size(x_39); +x_45 = lean_usize_of_nat(x_44); +lean_dec(x_44); +lean_inc(x_4); +lean_inc(x_3); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__8(x_45, x_8, x_39, x_3, x_4, x_43); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12; +lean_inc(x_4); +lean_inc(x_3); +x_50 = l_Lean_Compiler_checkpoint(x_49, x_47, x_22, x_3, x_4, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +lean_inc(x_47); +x_52 = l_Lean_Compiler_saveStage1Decls(x_47, x_3, x_4, x_51); +x_53 = !lean_is_exclusive(x_52); +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_52, 1); +x_55 = lean_ctor_get(x_52, 0); +lean_dec(x_55); +x_56 = lean_array_get_size(x_47); +x_57 = lean_nat_dec_lt(x_26, x_56); +if (x_57 == 0) +{ +lean_dec(x_56); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_52, 0, x_47); +return x_52; +} +else +{ +uint8_t x_58; +x_58 = lean_nat_dec_le(x_56, x_56); +if (x_58 == 0) +{ +lean_dec(x_56); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_52, 0, x_47); +return x_52; +} +else +{ +size_t x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_52); +x_59 = lean_usize_of_nat(x_56); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9(x_47, x_8, x_59, x_60, x_3, x_4, x_54); +if (lean_obj_tag(x_61) == 0) +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_61, 0); +lean_dec(x_63); +lean_ctor_set(x_61, 0, x_47); +return x_61; +} +else +{ +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_47); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +else +{ +uint8_t x_66; +lean_dec(x_47); +x_66 = !lean_is_exclusive(x_61); +if (x_66 == 0) +{ +return x_61; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_61, 0); +x_68 = lean_ctor_get(x_61, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_61); +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; +} +} +} +} +} +else +{ +lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_70 = lean_ctor_get(x_52, 1); +lean_inc(x_70); +lean_dec(x_52); +x_71 = lean_array_get_size(x_47); +x_72 = lean_nat_dec_lt(x_26, x_71); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_71); +lean_dec(x_4); +lean_dec(x_3); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_47); +lean_ctor_set(x_73, 1, x_70); +return x_73; +} +else +{ +uint8_t x_74; +x_74 = lean_nat_dec_le(x_71, x_71); +if (x_74 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* 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 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__6; -lean_inc(x_4); -lean_inc(x_3); -x_31 = l_Lean_Compiler_checkpoint(x_30, x_28, x_22, x_3, x_4, x_29); -if (lean_obj_tag(x_31) == 0) +lean_object* x_75; +lean_dec(x_71); +lean_dec(x_4); +lean_dec(x_3); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_47); +lean_ctor_set(x_75, 1, x_70); +return x_75; +} +else { -lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_array_get_size(x_28); -x_34 = lean_usize_of_nat(x_33); -lean_dec(x_33); -lean_inc(x_4); -lean_inc(x_3); -x_35 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4(x_34, x_8, x_28, x_3, x_4, x_32); -if (lean_obj_tag(x_35) == 0) +size_t x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_usize_of_nat(x_71); +lean_dec(x_71); +x_77 = lean_box(0); +x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9(x_47, x_8, x_76, x_77, x_3, x_4, x_70); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_80 = x_78; +} else { + lean_dec_ref(x_78); + x_80 = lean_box(0); +} +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_80; +} +lean_ctor_set(x_81, 0, x_47); +lean_ctor_set(x_81, 1, x_79); +return x_81; +} +else { -lean_object* x_36; lean_object* x_37; lean_object* 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 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8; -lean_inc(x_4); -lean_inc(x_3); -x_39 = l_Lean_Compiler_checkpoint(x_38, x_36, x_22, x_3, x_4, x_37); -if (lean_obj_tag(x_39) == 0) +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_47); +x_82 = lean_ctor_get(x_78, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_78, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_84 = x_78; +} else { + lean_dec_ref(x_78); + x_84 = lean_box(0); +} +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; +} +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} +} +} +} +} +else { -lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -lean_inc(x_36); -x_41 = l_Lean_Compiler_saveStage1Decls(x_36, x_3, x_4, x_40); +uint8_t x_86; +lean_dec(x_47); lean_dec(x_4); lean_dec(x_3); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +x_86 = !lean_is_exclusive(x_50); +if (x_86 == 0) { -lean_object* x_43; -x_43 = lean_ctor_get(x_41, 0); -lean_dec(x_43); -lean_ctor_set(x_41, 0, x_36); -return x_41; +return x_50; } else { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_dec(x_41); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_36); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_50, 0); +x_88 = lean_ctor_get(x_50, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_50); +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; +} } } else { -uint8_t x_46; -lean_dec(x_36); +uint8_t x_90; lean_dec(x_4); lean_dec(x_3); -x_46 = !lean_is_exclusive(x_39); -if (x_46 == 0) +x_90 = !lean_is_exclusive(x_46); +if (x_90 == 0) { -return x_39; +return x_46; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_39, 0); -x_48 = lean_ctor_get(x_39, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_39); -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_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_46, 0); +x_92 = lean_ctor_get(x_46, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_46); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } } else { -uint8_t x_50; +uint8_t x_94; +lean_dec(x_39); lean_dec(x_4); lean_dec(x_3); -x_50 = !lean_is_exclusive(x_35); -if (x_50 == 0) +x_94 = !lean_is_exclusive(x_42); +if (x_94 == 0) { -return x_35; +return x_42; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_35, 0); -x_52 = lean_ctor_get(x_35, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_35); -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; +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_42, 0); +x_96 = lean_ctor_get(x_42, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_42); +x_97 = lean_alloc_ctor(1, 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_54; -lean_dec(x_28); +uint8_t x_98; lean_dec(x_4); lean_dec(x_3); -x_54 = !lean_is_exclusive(x_31); -if (x_54 == 0) +x_98 = !lean_is_exclusive(x_38); +if (x_98 == 0) { -return x_31; +return x_38; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_38, 0); +x_100 = lean_ctor_get(x_38, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_38); +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_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_31, 0); -x_56 = lean_ctor_get(x_31, 1); -lean_inc(x_56); -lean_inc(x_55); +uint8_t x_102; lean_dec(x_31); -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; +lean_dec(x_4); +lean_dec(x_3); +x_102 = !lean_is_exclusive(x_34); +if (x_102 == 0) +{ +return x_34; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_34, 0); +x_104 = lean_ctor_get(x_34, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_34); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +return x_105; } } } else { -uint8_t x_58; +uint8_t x_106; lean_dec(x_4); lean_dec(x_3); -x_58 = !lean_is_exclusive(x_27); -if (x_58 == 0) +x_106 = !lean_is_exclusive(x_30); +if (x_106 == 0) { -return x_27; +return x_30; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_27, 0); -x_60 = lean_ctor_get(x_27, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_27); -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_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_30, 0); +x_108 = lean_ctor_get(x_30, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_30); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; +} } } } else { -uint8_t x_62; +uint8_t x_120; lean_dec(x_19); lean_dec(x_4); lean_dec(x_3); -x_62 = !lean_is_exclusive(x_23); -if (x_62 == 0) +x_120 = !lean_is_exclusive(x_23); +if (x_120 == 0) { return x_23; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_23, 0); -x_64 = lean_ctor_get(x_23, 1); -lean_inc(x_64); -lean_inc(x_63); +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_23, 0); +x_122 = lean_ctor_get(x_23, 1); +lean_inc(x_122); +lean_inc(x_121); lean_dec(x_23); -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; +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_66; +uint8_t x_124; lean_dec(x_4); lean_dec(x_3); -x_66 = !lean_is_exclusive(x_18); -if (x_66 == 0) +x_124 = !lean_is_exclusive(x_18); +if (x_124 == 0) { return x_18; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_18, 0); -x_68 = lean_ctor_get(x_18, 1); -lean_inc(x_68); -lean_inc(x_67); +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_18, 0); +x_126 = lean_ctor_get(x_18, 1); +lean_inc(x_126); +lean_inc(x_125); lean_dec(x_18); -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; +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_70; +uint8_t x_128; lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); -x_70 = !lean_is_exclusive(x_14); -if (x_70 == 0) +x_128 = !lean_is_exclusive(x_14); +if (x_128 == 0) { return x_14; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_14, 0); -x_72 = lean_ctor_get(x_14, 1); -lean_inc(x_72); -lean_inc(x_71); +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_14, 0); +x_130 = lean_ctor_get(x_14, 1); +lean_inc(x_130); +lean_inc(x_129); lean_dec(x_14); -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; +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_74; +uint8_t x_132; lean_dec(x_4); lean_dec(x_3); -x_74 = !lean_is_exclusive(x_9); -if (x_74 == 0) +x_132 = !lean_is_exclusive(x_9); +if (x_132 == 0) { return x_9; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_9, 0); -x_76 = lean_ctor_get(x_9, 1); -lean_inc(x_76); -lean_inc(x_75); +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_9, 0); +x_134 = lean_ctor_get(x_9, 1); +lean_inc(x_134); +lean_inc(x_133); lean_dec(x_9); -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; +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } @@ -2353,7 +3354,7 @@ lean_dec(x_13); x_21 = l_Lean_Compiler_shouldGenerateCode___closed__9; lean_inc(x_3); lean_inc(x_2); -x_22 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__5(x_1, x_19, x_20, x_21, x_2, x_3, x_4); +x_22 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__11(x_1, x_19, x_20, x_21, x_2, x_3, x_4); lean_dec(x_1); if (lean_obj_tag(x_22) == 0) { @@ -2443,7 +3444,39 @@ x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2(x_7, return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___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_isTracingEnabledFor___at_Lean_Compiler_compileStage1Impl___spec__3(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_addTrace___at_Lean_Compiler_compileStage1Impl___spec__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_addTrace___at_Lean_Compiler_compileStage1Impl___spec__5(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_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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) { +_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_2); +lean_dec(x_2); +x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6(x_7, x_8, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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; size_t x_8; lean_object* x_9; @@ -2451,11 +3484,11 @@ x_7 = lean_unbox_usize(x_1); lean_dec(x_1); x_8 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__3(x_7, x_8, x_3, x_4, x_5, x_6); +x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7(x_7, x_8, x_3, x_4, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___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) { _start: { size_t x_7; size_t x_8; lean_object* x_9; @@ -2463,11 +3496,37 @@ x_7 = lean_unbox_usize(x_1); lean_dec(x_1); x_8 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4(x_7, x_8, x_3, x_4, x_5, x_6); +x_9 = l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__8(x_7, x_8, x_3, x_4, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___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) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___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) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___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) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -2475,7 +3534,7 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__5(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__11(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_10; } @@ -2593,7 +3652,7 @@ lean_dec(x_1); return x_7; } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -2603,27 +3662,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; -x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__3; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; -x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__5; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -2633,7 +3692,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728_(lean_object* x_1) { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__2; +x_2 = l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -2642,128 +3711,221 @@ x_3 = 0; x_4 = l_Lean_registerTraceClass(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { -lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); -x_6 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1; -x_7 = 1; -x_8 = l_Lean_registerTraceClass(x_6, x_7, x_5); -if (lean_obj_tag(x_8) == 0) +x_6 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2; +x_7 = l_Lean_registerTraceClass(x_6, x_3, x_5); +if (lean_obj_tag(x_7) == 0) { -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_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2; -x_11 = l_Lean_registerTraceClass(x_10, x_7, x_9); +lean_object* x_8; lean_object* x_9; uint8_t 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_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1; +x_10 = 1; +x_11 = l_Lean_registerTraceClass(x_9, x_10, x_8); 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, 1); lean_inc(x_12); lean_dec(x_11); -x_13 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3; -x_14 = l_Lean_registerTraceClass(x_13, x_7, x_12); +x_13 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2; +x_14 = l_Lean_registerTraceClass(x_13, x_10, 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, 1); lean_inc(x_15); lean_dec(x_14); -x_16 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4; -x_17 = l_Lean_registerTraceClass(x_16, x_7, x_15); +x_16 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3; +x_17 = l_Lean_registerTraceClass(x_16, x_10, 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, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4; +x_20 = l_Lean_registerTraceClass(x_19, x_10, 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, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5; +x_23 = l_Lean_registerTraceClass(x_22, x_10, 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, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2; +x_26 = l_Lean_registerTraceClass(x_25, x_3, x_24); +return x_26; +} +else +{ +uint8_t x_27; +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 +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_20); +if (x_31 == 0) +{ +return x_20; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_20, 0); +x_33 = lean_ctor_get(x_20, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_20); +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; +} +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_17); +if (x_35 == 0) +{ return x_17; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_17, 0); +x_37 = lean_ctor_get(x_17, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_17); +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; +} +} +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_14); +if (x_39 == 0) { return x_14; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -x_20 = lean_ctor_get(x_14, 1); -lean_inc(x_20); -lean_inc(x_19); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +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_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; +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 { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_11); -if (x_22 == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_11); +if (x_43 == 0) { return x_11; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_11, 0); -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_inc(x_23); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +lean_inc(x_45); +lean_inc(x_44); lean_dec(x_11); -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; +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 { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_8); -if (x_26 == 0) +uint8_t x_47; +x_47 = !lean_is_exclusive(x_7); +if (x_47 == 0) { -return x_8; +return x_7; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_8, 0); -x_28 = lean_ctor_get(x_8, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_8); -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_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_7, 0); +x_49 = lean_ctor_get(x_7, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_7); +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 { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_4); -if (x_30 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_4); +if (x_51 == 0) { return x_4; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_4, 0); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); -lean_inc(x_31); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_4, 0); +x_53 = lean_ctor_get(x_4, 1); +lean_inc(x_53); +lean_inc(x_52); lean_dec(x_4); -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; +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } } @@ -2774,6 +3936,7 @@ lean_object* initialize_Lean_Compiler_TerminalCases(uint8_t builtin, lean_object lean_object* initialize_Lean_Compiler_CSE(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_Stage1(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_PullLocalDecls(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_Main(uint8_t builtin, lean_object* w) { lean_object * res; @@ -2797,6 +3960,9 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_Simp(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Compiler_PullLocalDecls(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Compiler_shouldGenerateCode___lambda__2___closed__1 = _init_l_Lean_Compiler_shouldGenerateCode___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Compiler_shouldGenerateCode___lambda__2___closed__1); l_Lean_Compiler_shouldGenerateCode___closed__1 = _init_l_Lean_Compiler_shouldGenerateCode___closed__1(); @@ -2879,8 +4045,24 @@ l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__15 = lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_checkpoint___spec__3___closed__15); l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__2___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__4___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__6___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Compiler_compileStage1Impl___spec__7___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__2); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__3); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__9___closed__4); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__2); +l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Compiler_compileStage1Impl___spec__10___closed__3); l_Lean_Compiler_compileStage1Impl___lambda__1___closed__1 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__1); l_Lean_Compiler_compileStage1Impl___lambda__1___closed__2 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__2(); @@ -2897,17 +4079,27 @@ l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7 = _init_l_Lean_Compile lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__7); l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8(); lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__8); +l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__9); +l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__10); +l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__11); +l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12 = _init_l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Compiler_compileStage1Impl___lambda__1___closed__12); l_Lean_Compiler_compile___closed__1 = _init_l_Lean_Compiler_compile___closed__1(); lean_mark_persistent(l_Lean_Compiler_compile___closed__1); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__1); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__2); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__3); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728____closed__4); -res = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_728_(lean_io_mk_world()); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__1); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__2); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__3); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__4); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989____closed__5); +res = l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_989_(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/PullLocalDecls.c b/stage0/stdlib/Lean/Compiler/PullLocalDecls.c new file mode 100644 index 000000000000..b06e602c15b5 --- /dev/null +++ b/stage0/stdlib/Lean/Compiler/PullLocalDecls.c @@ -0,0 +1,3131 @@ +// Lean compiler output +// Module: Lean.Compiler.PullLocalDecls +// Imports: Init Lean.Compiler.CompilerM Lean.Compiler.Decl +#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_Compiler_PullLocalDecls_dependsOn___boxed(lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Compiler_Decl_pullInstances___closed__1; +lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1___boxed(lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___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*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* l_Lean_Compiler_mkLetUsingScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3; +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* l_Array_append___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3; +static lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullLocalDecls___lambda__1(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_Compiler_Decl_mapValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_array_push(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1; +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___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_EXPORT lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_foldlMUnsafe_fold___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__5(lean_object*, size_t, size_t, lean_object*); +extern lean_object* l_Lean_levelZero; +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___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_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4___lambda__1___boxed(lean_object**); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); +lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_value(lean_object*); +extern lean_object* l_Lean_inheritedTraceOptions; +lean_object* lean_array_fget(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6; +lean_object* l_Lean_Compiler_mkLetDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___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*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7; +static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1(lean_object*, lean_object*); +lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabited___rarg(lean_object*, lean_object*); +lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitCases___closed__1; +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2; +lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1; +static lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2; +extern lean_object* l_Lean_instInhabitedExpr; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___boxed(lean_object*, lean_object*, 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*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__8; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4; +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5; +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isLambda(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_State_toPull___default; +lean_object* l_Lean_LocalDecl_fvarId(lean_object*); +lean_object* l_Lean_Compiler_visitLambdaCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_PullLocalDecls___hyg_1426_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7; +lean_object* l_Lean_LocalDecl_type(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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_Compiler_Decl_pullLocalDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___boxed(lean_object**); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Core_instMonadCoreM; +uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4; +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1; +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___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*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3(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_st_ref_set(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isFVar(lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4; +lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1; +static lean_object* l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_RBNode_insert___at_Lean_Compiler_check_checkBlock___spec__6(lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLambda(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3; +lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__3___boxed(lean_object**); +static lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6; +lean_object* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2; +lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_local_ctx_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4___boxed(lean_object**); +lean_object* l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Compiler_PullLocalDecls_dependsOn(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9; +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1; +extern lean_object* l_instInhabitedPUnit; +static lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_State_toPull___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_Lean_Expr_isFVar(x_2); +if (x_3 == 0) +{ +uint8_t x_4; +lean_dec(x_2); +x_4 = 0; +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_Expr_fvarId_x21(x_2); +x_6 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_1, x_5); +lean_dec(x_5); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = 0; +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_6); +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_Compiler_PullLocalDecls_dependsOn(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_alloc_closure((void*)(l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1___boxed), 2, 1); +lean_closure_set(x_3, 0, x_2); +x_4 = l_Lean_Expr_FindImpl_findUnsafe_x3f(x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +uint8_t x_6; +lean_dec(x_4); +x_6 = 1; +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Compiler_PullLocalDecls_dependsOn___lambda__1(x_1, x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_dependsOn___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Compiler_PullLocalDecls_dependsOn(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___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; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_4, 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; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_local_ctx_find(x_13, x_1); +lean_ctor_set(x_10, 0, x_14); +return x_10; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +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_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_local_ctx_find(x_17, x_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); +return x_19; +} +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_instMonadCoreM; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___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; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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) { +_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; +x_13 = lean_array_push(x_5, x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_box(0); +x_16 = l_Std_RBNode_insert___at_Lean_Compiler_check_checkBlock___spec__6(x_3, x_2, 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_14); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, 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; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.PullLocalDecls", 28); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.PullLocalDecls.mkLambda", 37); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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; lean_object* x_6; +x_1 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1; +x_2 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2; +x_3 = lean_unsigned_to_nat(31u); +x_4 = lean_unsigned_to_nat(70u); +x_5 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Init.Util", 9); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("getElem!", 8); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("index out of bounds", 19); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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; lean_object* x_6; +x_1 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5; +x_2 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6; +x_3 = lean_unsigned_to_nat(73u); +x_4 = lean_unsigned_to_nat(36u); +x_5 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7; +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_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 0); +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_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_17 = x_12; +} else { + lean_dec_ref(x_12); + x_17 = lean_box(0); +} +x_18 = lean_nat_dec_lt(x_2, x_3); +lean_dec(x_3); +if (x_18 == 0) +{ +lean_object* x_96; lean_object* x_97; +lean_dec(x_5); +lean_dec(x_2); +x_96 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__8; +x_97 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_96); +x_19 = x_97; +goto block_95; +} +else +{ +lean_object* x_98; +x_98 = lean_array_fget(x_5, x_2); +lean_dec(x_2); +lean_dec(x_5); +x_19 = x_98; +goto block_95; +} +block_95: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_19); +x_20 = l_Lean_Expr_fvarId_x21(x_19); +x_21 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +lean_dec(x_4); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4; +x_25 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2(x_24, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_25) == 0) +{ +uint8_t x_26; +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; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +if (lean_is_scalar(x_17)) { + x_28 = lean_alloc_ctor(0, 2, 0); +} else { + x_28 = x_17; +} +lean_ctor_set(x_28, 0, x_15); +lean_ctor_set(x_28, 1, x_16); +if (lean_is_scalar(x_14)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_14; +} +lean_ctor_set(x_29, 0, x_13); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_25, 0, x_30); +return x_25; +} +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_25, 1); +lean_inc(x_31); +lean_dec(x_25); +if (lean_is_scalar(x_17)) { + x_32 = lean_alloc_ctor(0, 2, 0); +} else { + x_32 = x_17; +} +lean_ctor_set(x_32, 0, x_15); +lean_ctor_set(x_32, 1, x_16); +if (lean_is_scalar(x_14)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_14; +} +lean_ctor_set(x_33, 0, x_13); +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_31); +return x_35; +} +} +else +{ +uint8_t x_36; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +x_36 = !lean_is_exclusive(x_25); +if (x_36 == 0) +{ +return x_25; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_25, 0); +x_38 = lean_ctor_get(x_25, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_25); +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; +x_40 = lean_ctor_get(x_22, 0); +lean_inc(x_40); +lean_dec(x_22); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_19); +lean_dec(x_4); +x_41 = lean_ctor_get(x_21, 1); +lean_inc(x_41); +lean_dec(x_21); +x_42 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4; +x_43 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2(x_42, x_6, x_7, x_8, x_9, x_10, x_41); +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; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +if (lean_is_scalar(x_17)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_17; +} +lean_ctor_set(x_46, 0, x_15); +lean_ctor_set(x_46, 1, x_16); +if (lean_is_scalar(x_14)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_14; +} +lean_ctor_set(x_47, 0, x_13); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_43, 0, x_48); +return x_43; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +if (lean_is_scalar(x_17)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_17; +} +lean_ctor_set(x_50, 0, x_15); +lean_ctor_set(x_50, 1, x_16); +if (lean_is_scalar(x_14)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_14; +} +lean_ctor_set(x_51, 0, x_13); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, 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); +return x_53; +} +} +else +{ +uint8_t x_54; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +x_54 = !lean_is_exclusive(x_43); +if (x_54 == 0) +{ +return x_43; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_43, 0); +x_56 = lean_ctor_get(x_43, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_43); +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; +x_58 = !lean_is_exclusive(x_21); +if (x_58 == 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; +x_59 = lean_ctor_get(x_21, 1); +x_60 = lean_ctor_get(x_21, 0); +lean_dec(x_60); +x_61 = lean_ctor_get(x_40, 1); +lean_inc(x_61); +x_62 = lean_ctor_get(x_40, 3); +lean_inc(x_62); +x_63 = lean_ctor_get(x_40, 4); +lean_inc(x_63); +lean_dec(x_40); +x_64 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_4, x_61); +lean_dec(x_4); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_63); +lean_dec(x_62); +lean_free_object(x_21); +lean_dec(x_17); +lean_dec(x_14); +x_65 = lean_box(0); +x_66 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_61, x_13, x_15, x_16, x_65, x_6, x_7, x_8, x_9, x_10, x_59); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_66; +} +else +{ +uint8_t x_67; +lean_dec(x_64); +lean_inc(x_13); +x_67 = l_Lean_Compiler_PullLocalDecls_dependsOn(x_63, x_13); +if (x_67 == 0) +{ +uint8_t x_68; +lean_inc(x_13); +x_68 = l_Lean_Compiler_PullLocalDecls_dependsOn(x_62, x_13); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_61); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_69 = lean_array_push(x_15, x_19); +if (lean_is_scalar(x_17)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_17; +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_16); +if (lean_is_scalar(x_14)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_14; +} +lean_ctor_set(x_71, 0, x_13); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_21, 0, x_72); +return x_21; +} +else +{ +lean_object* x_73; lean_object* x_74; +lean_free_object(x_21); +lean_dec(x_17); +lean_dec(x_14); +x_73 = lean_box(0); +x_74 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_61, x_13, x_15, x_16, x_73, x_6, x_7, x_8, x_9, x_10, x_59); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_74; +} +} +else +{ +lean_object* x_75; lean_object* x_76; +lean_dec(x_62); +lean_free_object(x_21); +lean_dec(x_17); +lean_dec(x_14); +x_75 = lean_box(0); +x_76 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_61, x_13, x_15, x_16, x_75, x_6, x_7, x_8, x_9, x_10, x_59); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_76; +} +} +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_21, 1); +lean_inc(x_77); +lean_dec(x_21); +x_78 = lean_ctor_get(x_40, 1); +lean_inc(x_78); +x_79 = lean_ctor_get(x_40, 3); +lean_inc(x_79); +x_80 = lean_ctor_get(x_40, 4); +lean_inc(x_80); +lean_dec(x_40); +x_81 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_4, x_78); +lean_dec(x_4); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_80); +lean_dec(x_79); +lean_dec(x_17); +lean_dec(x_14); +x_82 = lean_box(0); +x_83 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_78, x_13, x_15, x_16, x_82, x_6, x_7, x_8, x_9, x_10, x_77); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_83; +} +else +{ +uint8_t x_84; +lean_dec(x_81); +lean_inc(x_13); +x_84 = l_Lean_Compiler_PullLocalDecls_dependsOn(x_80, x_13); +if (x_84 == 0) +{ +uint8_t x_85; +lean_inc(x_13); +x_85 = l_Lean_Compiler_PullLocalDecls_dependsOn(x_79, x_13); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_78); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_86 = lean_array_push(x_15, x_19); +if (lean_is_scalar(x_17)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_17; +} +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_16); +if (lean_is_scalar(x_14)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_14; +} +lean_ctor_set(x_88, 0, x_13); +lean_ctor_set(x_88, 1, x_87); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, 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_77); +return x_90; +} +else +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_17); +lean_dec(x_14); +x_91 = lean_box(0); +x_92 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_78, x_13, x_15, x_16, x_91, x_6, x_7, x_8, x_9, x_10, x_77); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_92; +} +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_79); +lean_dec(x_17); +lean_dec(x_14); +x_93 = lean_box(0); +x_94 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__1(x_19, x_78, x_13, x_15, x_16, x_93, x_6, x_7, x_8, x_9, x_10, x_77); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_94; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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, lean_object* x_17, lean_object* x_18) { +_start: +{ +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* 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_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_apply_8(x_21, lean_box(0), x_19, x_13, x_14, x_15, x_16, x_17, x_18); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_12, 0); +lean_inc(x_23); +lean_dec(x_12); +x_24 = lean_nat_add(x_2, x_3); +lean_dec(x_2); +x_25 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24, x_11, x_3, x_23, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_10); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___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, 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) { +_start: +{ +uint8_t x_19; +x_19 = lean_nat_dec_le(x_10, x_9); +if (x_19 == 0) +{ +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; +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_8, x_22); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_7); +lean_inc(x_9); +x_25 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2), 11, 5); +lean_closure_set(x_25, 0, x_12); +lean_closure_set(x_25, 1, x_9); +lean_closure_set(x_25, 2, x_7); +lean_closure_set(x_25, 3, x_5); +lean_closure_set(x_25, 4, x_4); +x_26 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__3___boxed), 18, 11); +lean_closure_set(x_26, 0, x_1); +lean_closure_set(x_26, 1, x_9); +lean_closure_set(x_26, 2, x_11); +lean_closure_set(x_26, 3, x_2); +lean_closure_set(x_26, 4, x_3); +lean_closure_set(x_26, 5, x_4); +lean_closure_set(x_26, 6, x_5); +lean_closure_set(x_26, 7, x_6); +lean_closure_set(x_26, 8, x_7); +lean_closure_set(x_26, 9, x_23); +lean_closure_set(x_26, 10, x_10); +x_27 = lean_apply_10(x_24, lean_box(0), lean_box(0), x_25, x_26, x_13, x_14, x_15, x_16, x_17, x_18); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +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_4); +lean_dec(x_3); +lean_dec(x_2); +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +lean_dec(x_1); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_apply_8(x_29, lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +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_4); +lean_dec(x_3); +lean_dec(x_2); +x_31 = lean_ctor_get(x_1, 0); +lean_inc(x_31); +lean_dec(x_1); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_apply_8(x_32, lean_box(0), x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_33; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___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, 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_11) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* 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_3); +lean_dec(x_2); +x_18 = lean_ctor_get(x_11, 0); +lean_inc(x_18); +lean_dec(x_11); +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_apply_8(x_20, lean_box(0), x_18, x_12, x_13, x_14, x_15, x_16, x_17); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_11, 0); +lean_inc(x_22); +lean_dec(x_11); +x_23 = lean_nat_add(x_2, x_3); +lean_dec(x_2); +x_24 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_23, x_10, x_3, x_22, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_9); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___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, 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_le(x_9, x_8); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_eq(x_7, x_19); +if (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; +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_sub(x_7, x_21); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_6); +lean_inc(x_8); +x_24 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2), 11, 5); +lean_closure_set(x_24, 0, x_11); +lean_closure_set(x_24, 1, x_8); +lean_closure_set(x_24, 2, x_6); +lean_closure_set(x_24, 3, x_5); +lean_closure_set(x_24, 4, x_4); +x_25 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4___lambda__1___boxed), 17, 10); +lean_closure_set(x_25, 0, x_1); +lean_closure_set(x_25, 1, x_8); +lean_closure_set(x_25, 2, x_10); +lean_closure_set(x_25, 3, x_2); +lean_closure_set(x_25, 4, x_3); +lean_closure_set(x_25, 5, x_4); +lean_closure_set(x_25, 6, x_5); +lean_closure_set(x_25, 7, x_6); +lean_closure_set(x_25, 8, x_22); +lean_closure_set(x_25, 9, x_9); +x_26 = lean_apply_10(x_23, lean_box(0), lean_box(0), x_24, x_25, x_12, x_13, x_14, x_15, x_16, x_17); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_apply_8(x_28, lean_box(0), x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = lean_ctor_get(x_1, 0); +lean_inc(x_30); +lean_dec(x_1); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_apply_8(x_31, lean_box(0), x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_32; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__5(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; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Lean_Expr_fvarId_x21(x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_box(0); +x_11 = l_Std_RBNode_insert___at_Lean_Compiler_check_checkBlock___spec__6(x_4, x_7, x_10); +x_2 = x_9; +x_4 = x_11; +goto _start; +} +else +{ +return x_4; +} +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_ReaderT_instMonadLiftReaderT), 3, 2); +lean_closure_set(x_1, 0, lean_box(0)); +lean_closure_set(x_1, 1, lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_StateRefT_x27_lift), 4, 3); +lean_closure_set(x_1, 0, lean_box(0)); +lean_closure_set(x_1, 1, lean_box(0)); +lean_closure_set(x_1, 2, lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4() { +_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_Lean_Compiler_PullLocalDecls_mkLambda(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_10 = lean_st_ref_get(x_8, x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_st_ref_get(x_6, 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, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_8, x_14); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_get(x_5, 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); +x_21 = lean_box(0); +x_22 = lean_array_get_size(x_1); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_lt(x_23, x_22); +x_25 = lean_array_get_size(x_15); +x_26 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4; +lean_inc(x_1); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_1); +if (x_24 == 0) +{ +lean_dec(x_22); +lean_dec(x_1); +x_28 = x_21; +goto block_65; +} +else +{ +uint8_t x_66; +x_66 = lean_nat_dec_le(x_22, x_22); +if (x_66 == 0) +{ +lean_dec(x_22); +lean_dec(x_1); +x_28 = x_21; +goto block_65; +} +else +{ +size_t x_67; size_t x_68; lean_object* x_69; +x_67 = 0; +x_68 = lean_usize_of_nat(x_22); +lean_dec(x_22); +x_69 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__5(x_1, x_67, x_68, x_21); +lean_dec(x_1); +x_28 = x_69; +goto block_65; +} +} +block_65: +{ +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_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 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1; +x_31 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2; +x_32 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3; +x_33 = lean_unsigned_to_nat(1u); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc_n(x_25, 2); +x_34 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4(x_30, x_31, x_32, x_15, x_19, x_25, x_25, x_2, x_25, x_33, x_29, x_4, x_5, x_6, x_7, x_8, x_20); +lean_dec(x_25); +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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +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_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_st_ref_get(x_8, x_37); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_st_ref_take(x_6, x_41); +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_is_exclusive(x_43); +if (x_45 == 0) +{ +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_43, 1); +x_47 = l_Array_shrink___rarg(x_46, x_2); +lean_dec(x_2); +x_48 = l_Array_append___rarg(x_47, x_38); +lean_ctor_set(x_43, 1, x_48); +x_49 = lean_st_ref_set(x_6, x_43, x_44); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_Lean_Compiler_mkLambda(x_39, x_3, x_6, x_7, x_8, x_50); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_39); +return x_51; +} +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; lean_object* x_59; lean_object* x_60; +x_52 = lean_ctor_get(x_43, 0); +x_53 = lean_ctor_get(x_43, 1); +x_54 = lean_ctor_get(x_43, 2); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_43); +x_55 = l_Array_shrink___rarg(x_53, x_2); +lean_dec(x_2); +x_56 = l_Array_append___rarg(x_55, x_38); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_54); +x_58 = lean_st_ref_set(x_6, x_57, x_44); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Compiler_mkLambda(x_39, x_3, x_6, x_7, x_8, x_59); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_39); +return x_60; +} +} +else +{ +uint8_t x_61; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +x_61 = !lean_is_exclusive(x_34); +if (x_61 == 0) +{ +return x_34; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_34, 0); +x_63 = lean_ctor_get(x_34, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_34); +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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__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_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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); +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_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__3___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]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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, x_17, x_18); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___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]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__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, x_17, x_18); +lean_dec(x_8); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4___lambda__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: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___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, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__4___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_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__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, x_17); +lean_dec(x_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__5___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_Lean_Compiler_PullLocalDecls_mkLambda___spec__5(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLambda(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; 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_9 = l_Lean_Compiler_visitLambdaCore(x_1, x_5, x_6, 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); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_st_ref_get(x_7, x_11); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_5, x_15); +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, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_array_get_size(x_19); +lean_dec(x_19); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_12); +x_21 = l_Lean_Compiler_PullLocalDecls_visitLet(x_13, x_12, x_3, x_4, x_5, x_6, x_7, x_18); +if (lean_obj_tag(x_21) == 0) +{ +if (x_2 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +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 = l_Lean_Compiler_PullLocalDecls_mkLambda(x_12, x_20, x_22, x_3, x_4, x_5, x_6, x_7, x_23); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_20); +lean_dec(x_4); +lean_dec(x_3); +x_25 = lean_ctor_get(x_21, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_dec(x_21); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_27 = l_Lean_Compiler_mkLetUsingScope(x_25, x_5, x_6, x_7, x_26); +if (lean_obj_tag(x_27) == 0) +{ +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 = l_Lean_Compiler_mkLambda(x_12, x_28, x_5, x_6, x_7, x_29); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_30; +} +else +{ +uint8_t x_31; +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_31 = !lean_is_exclusive(x_27); +if (x_31 == 0) +{ +return x_27; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_27, 0); +x_33 = lean_ctor_get(x_27, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_27); +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; +} +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_20); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_35 = !lean_is_exclusive(x_21); +if (x_35 == 0) +{ +return x_21; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_21, 0); +x_37 = lean_ctor_get(x_21, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_21); +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; +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___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; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_inheritedTraceOptions; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___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; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1; +x_11 = lean_st_ref_get(x_10, 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; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_5, 2); +x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_13, x_14, x_1); +lean_dec(x_13); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_5, 2); +x_20 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_17, x_19, x_1); +lean_dec(x_17); +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_18); +return x_22; +} +} +} +static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3; +x_3 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4; +x_4 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5; +x_5 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set(x_5, 4, x_4); +lean_ctor_set(x_5, 5, x_2); +lean_ctor_set(x_5, 6, x_3); +lean_ctor_set(x_5, 7, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___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; 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; uint8_t x_27; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_st_ref_get(x_7, 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); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_get(x_7, x_12); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_5, x_15); +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 = lean_ctor_get(x_6, 2); +x_21 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6; +lean_inc(x_20); +x_22 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_19); +lean_ctor_set(x_22, 3, x_20); +x_23 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_2); +x_24 = lean_st_ref_take(x_7, x_18); +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_is_exclusive(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_28 = lean_ctor_get(x_25, 3); +x_29 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4; +x_30 = 0; +x_31 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_23); +lean_ctor_set(x_31, 2, x_29); +lean_ctor_set_uint8(x_31, sizeof(void*)*3, x_30); +lean_inc(x_9); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_9); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Std_PersistentArray_push___rarg(x_28, x_32); +lean_ctor_set(x_25, 3, x_33); +x_34 = lean_st_ref_set(x_7, x_25, x_26); +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 = lean_box(0); +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 = lean_box(0); +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; 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; 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_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +x_43 = lean_ctor_get(x_25, 2); +x_44 = lean_ctor_get(x_25, 3); +x_45 = lean_ctor_get(x_25, 4); +x_46 = lean_ctor_get(x_25, 5); +x_47 = lean_ctor_get(x_25, 6); +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_25); +x_48 = l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4; +x_49 = 0; +x_50 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_50, 0, x_1); +lean_ctor_set(x_50, 1, x_23); +lean_ctor_set(x_50, 2, x_48); +lean_ctor_set_uint8(x_50, sizeof(void*)*3, x_49); +lean_inc(x_9); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_9); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Std_PersistentArray_push___rarg(x_44, x_51); +x_53 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_53, 0, x_41); +lean_ctor_set(x_53, 1, x_42); +lean_ctor_set(x_53, 2, x_43); +lean_ctor_set(x_53, 3, x_52); +lean_ctor_set(x_53, 4, x_45); +lean_ctor_set(x_53, 5, x_46); +lean_ctor_set(x_53, 6, x_47); +x_54 = lean_st_ref_set(x_7, x_53, x_26); +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_54); + x_56 = lean_box(0); +} +x_57 = lean_box(0); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_56; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_dec(x_4); +x_11 = lean_array_push(x_1, x_2); +x_12 = l_Lean_Compiler_PullLocalDecls_visitLet(x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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; 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_10 = lean_st_ref_get(x_8, x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_st_ref_take(x_5, 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_Lean_LocalDecl_fvarId(x_1); +x_16 = lean_box(0); +x_17 = l_Std_RBNode_insert___at_Lean_Compiler_check_checkBlock___spec__6(x_13, x_15, x_16); +x_18 = lean_st_ref_set(x_5, x_17, x_14); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_apply_7(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_19); +return x_20; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.PullLocalDecls.visitLet", 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2() { +_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_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1; +x_2 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1; +x_3 = lean_unsigned_to_nat(67u); +x_4 = lean_unsigned_to_nat(47u); +x_5 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Compiler", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("pullLocalDecls", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4; +x_2 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("candidate", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6; +x_2 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___lambda__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_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +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_7); +x_14 = l_Lean_Compiler_mkLetDecl(x_1, x_2, x_6, x_3, x_10, x_11, x_12, x_13); +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_15); +x_17 = l_Lean_Expr_fvarId_x21(x_15); +x_18 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__1(x_17, x_8, x_9, x_10, x_11, x_12, x_16); +x_19 = lean_ctor_get(x_18, 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_dec(x_15); +lean_dec(x_5); +lean_dec(x_4); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2; +x_22 = l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1(x_21, x_8, x_9, x_10, x_11, x_12, x_20); +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; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_st_ref_get(x_12, x_23); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_9, x_26); +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); +lean_inc(x_8); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_24); +x_30 = lean_apply_6(x_8, x_24, x_28, x_10, x_11, x_12, x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; uint8_t x_32; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_unbox(x_31); +lean_dec(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_24); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec(x_30); +x_34 = lean_box(0); +x_35 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__1(x_4, x_15, x_5, x_34, x_8, x_9, x_10, x_11, x_12, x_33); +return x_35; +} +else +{ +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_30, 1); +lean_inc(x_36); +lean_dec(x_30); +lean_inc(x_15); +x_37 = lean_alloc_closure((void*)(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__1), 10, 3); +lean_closure_set(x_37, 0, x_4); +lean_closure_set(x_37, 1, x_15); +lean_closure_set(x_37, 2, x_5); +x_38 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8; +x_39 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2(x_38, x_8, x_9, x_10, x_11, x_12, x_36); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_15); +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_box(0); +x_44 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__2(x_24, x_37, x_43, x_8, x_9, x_10, x_11, x_12, x_42); +lean_dec(x_24); +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; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_46, 0, x_15); +x_47 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10; +x_48 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3(x_38, x_49, x_8, x_9, x_10, x_11, x_12, x_45); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__2(x_24, x_37, x_51, x_8, x_9, x_10, x_11, x_12, x_52); +lean_dec(x_51); +lean_dec(x_24); +return x_53; +} +} +} +else +{ +uint8_t x_54; +lean_dec(x_24); +lean_dec(x_15); +lean_dec(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); +x_54 = !lean_is_exclusive(x_30); +if (x_54 == 0) +{ +return x_30; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_30, 0); +x_56 = lean_ctor_get(x_30, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_30); +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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet(lean_object* x_1, lean_object* x_2, 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) == 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; uint8_t x_16; +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_ctor_get(x_1, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_12); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_dec(x_1); +x_14 = lean_expr_instantiate_rev(x_10, x_2); +lean_dec(x_10); +x_15 = lean_expr_instantiate_rev(x_11, x_2); +lean_dec(x_11); +x_16 = l_Lean_Expr_isLambda(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3(x_9, x_14, x_13, x_2, x_12, x_15, x_17, x_3, x_4, x_5, x_6, x_7, x_8); +return x_18; +} +else +{ +uint8_t x_19; lean_object* x_20; +x_19 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_20 = l_Lean_Compiler_PullLocalDecls_visitLambda(x_15, x_19, x_3, x_4, x_5, x_6, x_7, x_8); +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, 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); +x_24 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3(x_9, x_14, x_13, x_2, x_12, x_21, x_23, x_3, x_4, x_5, x_6, x_7, x_22); +return x_24; +} +else +{ +uint8_t x_25; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_9); +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_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_expr_instantiate_rev(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_29); +x_30 = l_Lean_Compiler_isCasesApp_x3f(x_29, x_6, x_7, x_8); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_30, 0); +lean_dec(x_33); +lean_ctor_set(x_30, 0, x_29); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_29); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_30, 1); +lean_inc(x_36); +lean_dec(x_30); +x_37 = lean_ctor_get(x_31, 0); +lean_inc(x_37); +lean_dec(x_31); +x_38 = l_Lean_Compiler_PullLocalDecls_visitCases(x_37, x_29, x_3, x_4, x_5, x_6, x_7, x_36); +return x_38; +} +} +else +{ +uint8_t x_39; +lean_dec(x_29); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_39 = !lean_is_exclusive(x_30); +if (x_39 == 0) +{ +return x_30; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_30, 0); +x_41 = lean_ctor_get(x_30, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_30); +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_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___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) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_3, x_2); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_1, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_1, x_15); +lean_dec(x_1); +x_17 = lean_array_get_size(x_5); +x_18 = lean_nat_dec_lt(x_2, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_nat_add(x_2, x_4); +lean_dec(x_2); +x_1 = x_16; +x_2 = x_19; +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +x_21 = lean_array_fget(x_5, x_2); +x_22 = lean_box(0); +x_23 = lean_array_fset(x_5, x_2, x_22); +x_24 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_25 = l_Lean_Compiler_PullLocalDecls_visitLambda(x_21, x_24, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_array_fset(x_23, x_2, x_26); +x_29 = lean_nat_add(x_2, x_4); +lean_dec(x_2); +x_1 = x_16; +x_2 = x_29; +x_5 = x_28; +x_11 = x_27; +goto _start; +} +else +{ +uint8_t x_31; +lean_dec(x_23); +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_25); +if (x_31 == 0) +{ +return x_25; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_25, 0); +x_33 = lean_ctor_get(x_25, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_25); +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; +} +} +} +} +else +{ +lean_object* x_35; +lean_dec(x_10); +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_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_5); +lean_ctor_set(x_35, 1, x_11); +return x_35; +} +} +else +{ +lean_object* x_36; +lean_dec(x_10); +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_5); +lean_ctor_set(x_36, 1, x_11); +return x_36; +} +} +} +static lean_object* _init_l_Lean_Compiler_PullLocalDecls_visitCases___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitCases(lean_object* x_1, lean_object* x_2, 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; lean_object* x_19; lean_object* x_20; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_9); +x_11 = l_Lean_Compiler_PullLocalDecls_visitCases___closed__1; +lean_inc(x_10); +x_12 = lean_mk_array(x_10, x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_10, x_13); +lean_dec(x_10); +lean_inc(x_2); +x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_12, x_14); +x_16 = lean_ctor_get(x_1, 3); +lean_inc(x_16); +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_16, 2); +lean_inc(x_19); +lean_dec(x_16); +lean_inc(x_17); +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___spec__1(x_17, x_18, x_17, x_19, x_15, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_19); +lean_dec(x_17); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = l_Lean_Expr_getAppFn(x_2); +lean_dec(x_2); +x_24 = l_Lean_mkAppN(x_23, x_22); +lean_ctor_set(x_20, 0, x_24); +return x_20; +} +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_20, 0); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_20); +x_27 = l_Lean_Expr_getAppFn(x_2); +lean_dec(x_2); +x_28 = l_Lean_mkAppN(x_27, x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) +{ +return x_20; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_20); +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_Compiler_PullLocalDecls_visitLambda___boxed(lean_object* x_1, lean_object* x_2, 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_Compiler_PullLocalDecls_visitLambda(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___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: +{ +lean_object* x_8; +x_8 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2(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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___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_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___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); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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_Lean_Compiler_PullLocalDecls_visitLet___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_PullLocalDecls_visitLet___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) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3(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); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___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: +{ +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_visitCases___spec__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); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullLocalDecls___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; 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_7 = lean_box(0); +x_8 = lean_st_ref_get(x_5, x_6); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_mk_ref(x_7, x_9); +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; +lean_inc(x_5); +lean_inc(x_11); +x_14 = l_Lean_Compiler_PullLocalDecls_visitLambda(x_2, x_13, x_1, x_11, x_3, x_4, x_5, x_12); +if (lean_obj_tag(x_14) == 0) +{ +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_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_st_ref_get(x_5, x_16); +lean_dec(x_5); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_st_ref_get(x_11, x_18); +lean_dec(x_11); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_15); +return x_19; +} +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_15); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +lean_dec(x_11); +lean_dec(x_5); +x_24 = !lean_is_exclusive(x_14); +if (x_24 == 0) +{ +return x_14; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_14, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_14); +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_Lean_Compiler_Decl_pullLocalDecls(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_closure((void*)(l_Lean_Compiler_Decl_pullLocalDecls___lambda__1), 6, 1); +lean_closure_set(x_6, 0, x_2); +x_7 = l_Lean_Compiler_Decl_mapValue(x_1, x_6, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances___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; lean_object* x_9; +x_7 = l_Lean_LocalDecl_type(x_1); +x_8 = l_Lean_Compiler_isClass_x3f(x_7, x_4, x_5, x_6); +lean_dec(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = l_Lean_LocalDecl_value(x_1); +if (lean_obj_tag(x_12) == 11) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 2); +lean_inc(x_13); +lean_dec(x_12); +if (lean_obj_tag(x_13) == 1) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_2, x_14); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; lean_object* x_17; +x_16 = 0; +x_17 = lean_box(x_16); +lean_ctor_set(x_8, 0, x_17); +return x_8; +} +else +{ +uint8_t x_18; lean_object* x_19; +lean_dec(x_15); +x_18 = 1; +x_19 = lean_box(x_18); +lean_ctor_set(x_8, 0, x_19); +return x_8; +} +} +else +{ +uint8_t x_20; lean_object* x_21; +lean_dec(x_13); +x_20 = 0; +x_21 = lean_box(x_20); +lean_ctor_set(x_8, 0, x_21); +return x_8; +} +} +else +{ +uint8_t x_22; lean_object* x_23; +lean_dec(x_12); +x_22 = 0; +x_23 = lean_box(x_22); +lean_ctor_set(x_8, 0, x_23); +return x_8; +} +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_8, 1); +lean_inc(x_24); +lean_dec(x_8); +x_25 = l_Lean_LocalDecl_value(x_1); +if (lean_obj_tag(x_25) == 11) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_25, 2); +lean_inc(x_26); +lean_dec(x_25); +if (lean_obj_tag(x_26) == 1) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l_Std_RBNode_findCore___at_Lean_Compiler_check_checkBlock___spec__3(x_2, x_27); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_29 = 0; +x_30 = lean_box(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_24); +return x_31; +} +else +{ +uint8_t x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_28); +x_32 = 1; +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_24); +return x_34; +} +} +else +{ +uint8_t x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_26); +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); +return x_37; +} +} +else +{ +uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_25); +x_38 = 0; +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); +return x_40; +} +} +} +else +{ +uint8_t x_41; +lean_dec(x_9); +x_41 = !lean_is_exclusive(x_8); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_8, 0); +lean_dec(x_42); +x_43 = 1; +x_44 = lean_box(x_43); +lean_ctor_set(x_8, 0, x_44); +return x_8; +} +else +{ +lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_dec(x_8); +x_46 = 1; +x_47 = lean_box(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_45); +return x_48; +} +} +} +} +static lean_object* _init_l_Lean_Compiler_Decl_pullInstances___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Decl_pullInstances___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances(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_Compiler_Decl_pullInstances___closed__1; +x_6 = l_Lean_Compiler_Decl_pullLocalDecls(x_1, x_5, x_2, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_pullInstances___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_Compiler_Decl_pullInstances___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_Lean_Compiler_initFn____x40_Lean_Compiler_PullLocalDecls___hyg_1426_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_2 = l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8; +x_3 = 0; +x_4 = l_Lean_registerTraceClass(x_2, x_3, x_1); +return x_4; +} +} +lean_object* initialize_Init(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_CompilerM(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Compiler_Decl(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Compiler_PullLocalDecls(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(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_CompilerM(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Compiler_Decl(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Compiler_PullLocalDecls_State_toPull___default = _init_l_Lean_Compiler_PullLocalDecls_State_toPull___default(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_State_toPull___default); +l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__1); +l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__2); +l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__3); +l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__2___closed__4); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__1); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__2); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__3); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__4); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__5); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__6); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__7); +l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__8 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__8(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_PullLocalDecls_mkLambda___spec__3___lambda__2___closed__8); +l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1 = _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_mkLambda___closed__1); +l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2 = _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_mkLambda___closed__2); +l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3 = _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_mkLambda___closed__3); +l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4 = _init_l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_mkLambda___closed__4); +l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__1); +l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2 = _init_l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Compiler_PullLocalDecls_visitLet___spec__1___closed__2); +l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Compiler_PullLocalDecls_visitLet___spec__2___closed__1); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__1 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__1(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__1); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__2); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__3); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__4); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__5); +l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6 = _init_l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_PullLocalDecls_visitLet___spec__3___closed__6); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__1); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__2); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__3); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__4); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__5); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__6); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__7); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__8); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__9); +l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10 = _init_l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitLet___lambda__3___closed__10); +l_Lean_Compiler_PullLocalDecls_visitCases___closed__1 = _init_l_Lean_Compiler_PullLocalDecls_visitCases___closed__1(); +lean_mark_persistent(l_Lean_Compiler_PullLocalDecls_visitCases___closed__1); +l_Lean_Compiler_Decl_pullInstances___closed__1 = _init_l_Lean_Compiler_Decl_pullInstances___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_pullInstances___closed__1); +res = l_Lean_Compiler_initFn____x40_Lean_Compiler_PullLocalDecls___hyg_1426_(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/Compiler/Simp.c b/stage0/stdlib/Lean/Compiler/Simp.c index 7ae840909201..4dd2c56430e0 100644 --- a/stage0/stdlib/Lean/Compiler/Simp.c +++ b/stage0/stdlib/Lean/Compiler/Simp.c @@ -13,286 +13,357 @@ #ifdef __cplusplus extern "C" { #endif -lean_object* l_Lean_Compiler_Decl_ensureUniqueLetVarNames(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__5; -lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at_Lean_Compiler_Simp_OccInfo_add___spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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_Compiler_Decl_simp_x3f___closed__3; -lean_object* l_Lean_Compiler_visitLetImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2___boxed(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp___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_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Compiler_Simp_inlineCandidate_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* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInline___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___spec__4(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* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Compiler_InferType_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Compiler_Simp_visitLambda___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_Simp_OccInfo_map___default___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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_Std_Range_forIn_loop___at_Lean_Compiler_Simp_visitCases___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_Compiler_Simp_inlineApp___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___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_Compiler_Simp_visitCases___closed__4; uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Compiler_mkLetUsingScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Compiler_Decl_simp___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_instInhabitedState___closed__1; +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1; -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_instInhabitedState; static lean_object* l_Lean_Compiler_Simp_visitCases___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5; +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1; lean_object* l_Lean_Compiler_getStage1Decl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1; -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1; -LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_Context_localInline___default; -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__2; -static lean_object* l_Lean_Compiler_Simp_visitCases___closed__6; -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findCtor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__5; +static lean_object* l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpProj_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_simpProj_x3f___closed__1; static lean_object* l_Lean_Compiler_Simp_visitCases___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg___lambda__1(lean_object*); +static lean_object* l_Lean_Compiler_Simp_instReprOcc___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpAppApp_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_OccInfo_map___default; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Compiler_Decl_simp___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__11(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_Simp_collectInlineStats___closed__4; -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1; -static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2; +uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Expr_constructorApp_x3f(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3(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_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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*); +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpAppApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitValue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__12; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_InlineStats_format___spec__3(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3; +static lean_object* l_Lean_Compiler_Simp_simpProj_x3f___closed__3; lean_object* lean_array_push(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* lean_expr_lift_loose_bvars(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__2(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4; +lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Compiler_Decl_simp_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1; +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2; static lean_object* l_Lean_Compiler_Decl_simp_x3f___closed__2; -lean_object* l_Lean_Compiler_withNewScopeImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goLambda___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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_Compiler_Simp_visitCases___closed__3; -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__8; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5; -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Compiler_Simp_OccInfo_add___spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_Simp_OccInfo_map___default___spec__1(lean_object*); +static lean_object* l_Lean_Compiler_Simp_simpProj_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__7(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_Simp_collectInlineStats_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__2; -static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5; +uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_attachOptJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_HashMap_toList___at_Lean_Compiler_Simp_InlineStats_format___spec__1(lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5; -static lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3; +static lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_State_occInfo___default; extern lean_object* l_Lean_levelZero; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_numOccs___default; +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_visitLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___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_Compiler_Simp_OccInfo_add(lean_object*, lean_object*); +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3; +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Decl_simp_x3f___closed__5; -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6; -static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1; -lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); +size_t lean_uint64_to_usize(uint64_t); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__6; +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp___closed__6; +static lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3; +static lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_instReprOcc; +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__5; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l_Lean_Name_hash___override(lean_object*); lean_object* l_Lean_LocalDecl_value(lean_object*); +static lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__3; +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1; extern lean_object* l_Lean_inheritedTraceOptions; -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__5; +LEAN_EXPORT lean_object* l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4___boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f(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_Compiler_Simp_InlineStats_shouldInline___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___boxed(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_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_betaReduce(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkLetDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__6; -static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2; -LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_OccInfo_format(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2; +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__1; +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10; +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4; lean_object* l_Std_mkHashMapImp___rarg(lean_object*); +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3; +LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_instInhabitedOcc; lean_object* l_instInhabited___rarg(lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp___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_Compiler_Simp_inlineApp___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__6; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInline(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInlineLocal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_add___spec__7(lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Config_smallThreshold___default; lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); -static lean_object* l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__8; +static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2; +static lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Compiler_Decl_simp_x3f___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Compiler_Simp_Occ_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_mkFreshLetVarIdx(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_getLambdaArity(lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4; extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13; +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3; +size_t lean_usize_modn(size_t, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp___closed__4; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f___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_Compiler_Decl_simp_x3f___lambda__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_Expr_isConst(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_InlineStats_format___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f___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_EXPORT lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_size___default; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__7; +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_bvar___override(lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__3; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4; +LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_attachJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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_object*); +lean_object* l_Lean_Compiler_visitLambdaCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1; +lean_object* l_Lean_Compiler_onlyOneExitPoint(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1; +static lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findLambda_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_State_simplified___default; -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2; -static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats___closed__7; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Decl_simp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1; +lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_instInhabitedOccInfo; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpProj_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); uint8_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(lean_object*, uint8_t, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6; +LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__1(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f(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_Lean_Compiler_Decl_simp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_OccInfo_format___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3(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_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_instToFormatInlineStats; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_OccInfo_format___spec__3(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* l_Lean_Compiler_mkAuxLetDeclName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_beta(lean_object*, lean_object*); lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadCoreM; uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_format(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_Compiler_Decl_simp_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_Compiler_Simp_collectInlineStats_goValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___closed__7; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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* l_Lean_Compiler_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_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_Lean_Compiler_Simp_Occ_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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*); +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2; +uint8_t l_Lean_Expr_isConstructorApp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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*, lean_object*); +uint8_t l_Lean_LocalDecl_isJp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findLambda_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +static lean_object* l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7; +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isAnyType(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLet___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_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1(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___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11; uint8_t l_Lean_Expr_isFVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3___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*); +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1; lean_object* l_Lean_Compiler_Decl_getArity(lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__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_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Context_config___default; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_Compiler_Simp_collectInlineStats_goValue___closed__1; lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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*); static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1; -static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isProj(lean_object*); lean_object* l_Lean_Compiler_mkJpDeclIfNotSimple(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findExpr(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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*); -static lean_object* l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4; +static lean_object* l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1; +LEAN_EXPORT lean_object* l_Std_HashMap_toList___at_Lean_Compiler_Simp_OccInfo_format___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_visitCases___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* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___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_Compiler_Simp_Occ_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__2; -static lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___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* l_Lean_Compiler_mkLocalDecl(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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_Compiler_Simp_internalize_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInlineLocal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___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*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4; static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___closed__1; +lean_object* l_Lean_Compiler_ensureUniqueLetVarNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_OptionT_instMonadOptionT___rarg(lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__8; lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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*, lean_object*); +lean_object* lean_local_ctx_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLet___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_Simp_mkFlatLet_go(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_Compiler_getLCNFSize(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkAuxLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9; lean_object* l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Compiler_Simp_inlineCandidate_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_levelOne; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_toCtorIdx(uint8_t); static lean_object* l_Lean_Compiler_Simp_visitCases___closed__5; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_InlineStats_shouldInline(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_instToFormatOccInfo; +static lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296_(uint8_t, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +lean_object* l_Lean_Compiler_lcnfSizeLe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1; +static lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2; +static lean_object* l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___boxed(lean_object*); static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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_nat_to_int(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710_(lean_object*); -lean_object* l_Lean_Compiler_visitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_go___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456_(lean_object*); +static lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2; +lean_object* l_Lean_Compiler_isEmptyType(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____boxed(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Config_increaseFactor___default; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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*); -static lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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*); +uint8_t l_Lean_Expr_isLet(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___lambda__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___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__1; +static lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findLambda_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -303,7 +374,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_1, 0); lean_inc(x_6); lean_dec(x_1); -x_7 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_6, x_2, x_3, x_4, x_5); +x_7 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(x_6, x_2, x_3, x_4, x_5); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); if (lean_obj_tag(x_8) == 0) @@ -540,7 +611,7 @@ case 1: lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); -x_8 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointChecker_containsNoJp___spec__2(x_7, x_3, x_4, x_5, x_6); +x_8 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_declIsInScope___spec__1(x_7, x_3, x_4, x_5, x_6); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); if (lean_obj_tag(x_9) == 0) @@ -662,32 +733,359 @@ lean_dec(x_3); return x_8; } } -static lean_object* _init_l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_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_Lean_Compiler_Simp_Occ_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_Lean_Compiler_Simp_Occ_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_Occ_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_Occ_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Compiler_Simp_Occ_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_Occ_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_Lean_Compiler_Simp_Occ_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp.Occ.once", 27); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Std_mkHashMapImp___rarg(x_1); +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__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___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3; +x_2 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2; +x_3 = lean_alloc_ctor(3, 2, 0); +lean_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_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(5, 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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6; +x_2 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2; +x_3 = lean_alloc_ctor(3, 2, 0); +lean_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_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(5, 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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp.Occ.many", 27); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9; +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Compiler_Simp_InlineStats_numOccs___default() { +static lean_object* _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3; +x_2 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10; +x_3 = lean_alloc_ctor(3, 2, 0); +lean_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_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(5, 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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6; +x_2 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10; +x_3 = lean_alloc_ctor(3, 2, 0); +lean_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_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(5, 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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296_(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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____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_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_instReprOcc___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_instReprOcc() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1; +x_1 = l_Lean_Compiler_Simp_instReprOcc___closed__1; +return x_1; +} +} +static uint8_t _init_l_Lean_Compiler_Simp_instInhabitedOcc() { +_start: +{ +uint8_t x_1; +x_1 = 0; return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_InlineStats_size___default() { +LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_Simp_OccInfo_map___default___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_mkHashMapImp___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_OccInfo_map___default() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Std_mkHashMapImp___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Compiler_Simp_OccInfo_map___default___spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_mkHashMap___at_Lean_Compiler_Simp_OccInfo_map___default___spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Std_mkHashMapImp___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_instInhabitedOccInfo() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1; +x_1 = l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -714,7 +1112,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_InlineStats_format___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_OccInfo_format___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -723,7 +1121,7 @@ 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_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2(x_4, x_6); +x_7 = l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2(x_4, x_6); lean_dec(x_6); x_8 = 1; x_9 = lean_usize_add(x_2, x_8); @@ -737,7 +1135,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Std_HashMap_toList___at_Lean_Compiler_Simp_InlineStats_format___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_HashMap_toList___at_Lean_Compiler_Simp_OccInfo_format___spec__1(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; @@ -770,14 +1168,14 @@ size_t x_8; size_t x_9; lean_object* x_10; x_8 = 0; x_9 = lean_usize_of_nat(x_4); lean_dec(x_4); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_InlineStats_format___spec__3(x_3, x_8, x_9, x_2); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_OccInfo_format___spec__3(x_3, x_8, x_9, x_2); lean_dec(x_3); return x_10; } } } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__1() { _start: { lean_object* x_1; @@ -785,17 +1183,17 @@ x_1 = lean_mk_string_from_bytes("\n", 1); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1; +x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__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_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3() { _start: { lean_object* x_1; @@ -803,17 +1201,17 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3; +x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5() { _start: { lean_object* x_1; @@ -821,139 +1219,91 @@ x_1 = lean_mk_string_from_bytes(" ↦ ", 5); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5; +x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7() { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(", ", 2); -return x_1; -} -} -static lean_object* _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8() { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7; -x_2 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_2, 0, x_1); return x_2; } -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_dec(x_1); -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_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; 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; +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_2, 1); +lean_dec(x_1); +x_5 = lean_ctor_get(x_3, 0); lean_inc(x_5); -lean_dec(x_2); -x_6 = lean_ctor_get(x_4, 0); +x_6 = lean_ctor_get(x_3, 1); lean_inc(x_6); -x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -lean_dec(x_4); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_inc(x_6); -x_9 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_8, x_6); -if (lean_obj_tag(x_9) == 0) -{ -lean_dec(x_7); -lean_dec(x_6); -x_2 = x_5; -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; 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_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2; +lean_dec(x_3); +x_7 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2; +x_8 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +x_9 = 1; +x_10 = l_Lean_Name_toString(x_5, x_9); +x_11 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4; x_13 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_13, 0, x_3); -lean_ctor_set(x_13, 1, x_12); -x_14 = 1; -x_15 = l_Lean_Name_toString(x_6, x_14); -x_16 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_16, 0, x_15); -x_17 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4; -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 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6; +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6; +x_15 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_unbox(x_6); +lean_dec(x_6); +x_18 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296_(x_17, x_16); +x_19 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_18); x_20 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Nat_repr(x_7); -x_22 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8; -x_25 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Nat_repr(x_11); -x_27 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_28, 0, x_25); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_17); -x_30 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_30, 0, x_13); -lean_ctor_set(x_30, 1, x_29); -x_2 = x_5; -x_3 = x_30; +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_12); +x_21 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +x_1 = x_4; +x_2 = x_21; goto _start; } } } -} -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_format(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_OccInfo_format(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); -lean_inc(x_2); -x_3 = l_Std_HashMap_toList___at_Lean_Compiler_Simp_InlineStats_format___spec__1(x_2); -x_4 = lean_box(0); -x_5 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4(x_1, x_3, x_4); -return x_5; +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_HashMap_toList___at_Lean_Compiler_Simp_OccInfo_format___spec__1(x_1); +x_3 = lean_box(0); +x_4 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4(x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_InlineStats_format___spec__2(x_1, x_2); +x_3 = l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_format___spec__2(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_InlineStats_format___spec__3___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_Lean_Compiler_Simp_OccInfo_format___spec__3___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; @@ -961,5469 +1311,8263 @@ 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_Lean_Compiler_Simp_InlineStats_format___spec__3(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_Simp_OccInfo_format___spec__3(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1() { _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_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_4, x_2); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = 0; -return x_6; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_OccInfo_format), 1, 0); +return x_1; } -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 = lean_unsigned_to_nat(1u); -x_9 = lean_nat_dec_eq(x_7, x_8); -lean_dec(x_7); -return x_9; } +static lean_object* _init_l_Lean_Compiler_Simp_instToFormatOccInfo() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1; +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_Compiler_Simp_InlineStats_shouldInline(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__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_inc(x_2); -x_4 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_3, x_2); -if (lean_obj_tag(x_4) == 0) +if (lean_obj_tag(x_2) == 0) { -uint8_t x_5; -lean_dec(x_2); -lean_dec(x_1); -x_5 = 0; -return x_5; +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } 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 = lean_unsigned_to_nat(1u); -x_8 = lean_nat_dec_eq(x_6, x_7); -lean_dec(x_6); -if (x_8 == 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_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = lean_box(0); -x_10 = l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1(x_1, x_2, x_9); -return x_10; +x_2 = x_6; +goto _start; } else { -uint8_t x_11; -lean_dec(x_2); -lean_dec(x_1); -x_11 = 1; -return x_11; +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; } } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Compiler_Simp_InlineStats_shouldInline___lambda__1(x_1, x_2, x_3); +lean_object* x_3; lean_object* x_4; uint64_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_array_get_size(x_3); +x_5 = l_Lean_Name_hash___override(x_2); +x_6 = lean_uint64_to_usize(x_5); +x_7 = lean_usize_modn(x_6, x_4); +lean_dec(x_4); +x_8 = lean_array_uget(x_3, x_7); lean_dec(x_3); -x_5 = lean_box(x_4); -return x_5; +x_9 = l_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__2(x_2, x_8); +lean_dec(x_8); +lean_dec(x_2); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_InlineStats_shouldInline___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Compiler_Simp_InlineStats_shouldInline(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; +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_2, 2); +x_6 = lean_name_eq(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; } -static lean_object* _init_l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_InlineStats_format), 1, 0); -return x_1; +uint8_t x_8; +x_8 = 1; +return x_8; } } -static lean_object* _init_l_Lean_Compiler_Simp_instToFormatInlineStats() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1; -return x_1; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___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_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_add___spec__7(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_apply_5(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -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_apply_6(x_2, x_9, x_3, x_4, x_5, x_6, x_10); -return x_11; +return x_1; } else { -uint8_t x_12; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_8); -if (x_12 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -return x_8; +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Name_hash___override(x_4); +x_8 = lean_uint64_to_usize(x_7); +x_9 = lean_usize_modn(x_8, x_6); +lean_dec(x_6); +x_10 = lean_array_uget(x_1, x_9); +lean_ctor_set(x_2, 2, x_10); +x_11 = lean_array_uset(x_1, x_9, x_2); +x_1 = x_11; +x_2 = x_5; +goto _start; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_8, 0); -x_14 = lean_ctor_get(x_8, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* 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_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -lean_dec(x_8); -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_dec(x_2); +x_16 = lean_array_get_size(x_1); +x_17 = l_Lean_Name_hash___override(x_13); +x_18 = lean_uint64_to_usize(x_17); +x_19 = lean_usize_modn(x_18, x_16); +lean_dec(x_16); +x_20 = lean_array_uget(x_1, x_19); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_14); +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_15; +goto _start; } } } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at_Lean_Compiler_Simp_OccInfo_add___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1___rarg), 7, 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) +{ +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_AssocList_foldlM___at_Lean_Compiler_Simp_OccInfo_add___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_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_collectInlineStats_goLambda___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_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Compiler_Simp_OccInfo_add___spec__5(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; lean_object* x_8; -x_7 = lean_apply_1(x_1, x_2); -x_8 = l_Lean_Compiler_withNewScopeImp___rarg(x_7, x_3, x_4, x_5, x_6); -return x_8; +lean_object* x_3; lean_object* x_4; 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_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_mul(x_3, x_4); +lean_dec(x_3); +x_6 = lean_box(0); +x_7 = lean_mk_array(x_5, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Std_HashMapImp_moveEntries___at_Lean_Compiler_Simp_OccInfo_add___spec__6(x_8, x_2, x_7); +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_Compiler_Simp_collectInlineStats_goLambda___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_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(lean_object* x_1, uint8_t x_2, lean_object* x_3) { _start: { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_dec(x_1); -x_8 = l_Lean_Compiler_Simp_collectInlineStats_go(x_7, x_2, x_3, x_4, x_5, x_6); -return x_8; +x_4 = lean_box(0); +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; 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 = lean_name_eq(x_6, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_collectInlineStats_goLambda___lambda__1), 6, 0); -return x_1; +lean_object* x_11; +lean_dec(x_7); +lean_dec(x_6); +x_11 = lean_box(x_2); +lean_ctor_set(x_3, 1, x_11); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goLambda(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: +else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_7, 0, x_1); -x_8 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1; -x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1___rarg), 7, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__2(x_10, x_2, x_3, x_4, x_5, x_6); -return x_11; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_3, 1); +x_14 = lean_ctor_get(x_3, 2); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_3); +x_15 = lean_name_eq(x_12, x_1); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(x_1, x_2, x_14); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_12); +lean_ctor_set(x_17, 1, x_13); +lean_ctor_set(x_17, 2, x_16); +return x_17; } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Init.Util", 9); -return x_1; +lean_object* x_18; lean_object* x_19; +lean_dec(x_13); +lean_dec(x_12); +x_18 = lean_box(x_2); +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_14); +return x_19; +} +} } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2() { +} +LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3(lean_object* x_1, lean_object* x_2, uint8_t x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("getElem!", 8); +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; uint64_t x_8; size_t x_9; size_t x_10; lean_object* x_11; uint8_t x_12; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_array_get_size(x_6); +x_8 = l_Lean_Name_hash___override(x_2); +x_9 = lean_uint64_to_usize(x_8); +x_10 = lean_usize_modn(x_9, x_7); +x_11 = lean_array_uget(x_6, x_10); +x_12 = l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4(x_2, x_11); +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; uint8_t x_19; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_5, x_13); +lean_dec(x_5); +x_15 = lean_box(x_3); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_11); +x_17 = lean_array_uset(x_6, x_10, x_16); +x_18 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_14); +x_19 = lean_nat_dec_le(x_18, x_7); +lean_dec(x_7); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +lean_free_object(x_1); +x_20 = l_Std_HashMapImp_expand___at_Lean_Compiler_Simp_OccInfo_add___spec__5(x_14, x_17); +return x_20; +} +else +{ +lean_ctor_set(x_1, 1, x_17); +lean_ctor_set(x_1, 0, x_14); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("index out of bounds", 19); +lean_object* x_21; lean_object* x_22; +lean_dec(x_7); +x_21 = l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(x_2, x_3, x_11); +x_22 = lean_array_uset(x_6, x_10, x_21); +lean_ctor_set(x_1, 1, x_22); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4() { -_start: +else { -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_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(73u); -x_4 = lean_unsigned_to_nat(36u); -x_5 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, 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_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; size_t x_27; size_t x_28; lean_object* x_29; uint8_t x_30; +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_1); +x_25 = lean_array_get_size(x_24); +x_26 = l_Lean_Name_hash___override(x_2); +x_27 = lean_uint64_to_usize(x_26); +x_28 = lean_usize_modn(x_27, x_25); +x_29 = lean_array_uget(x_24, x_28); +x_30 = l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4(x_2, x_29); +if (x_30 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 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; uint8_t x_37; +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_add(x_23, x_31); +lean_dec(x_23); +x_33 = lean_box(x_3); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_2); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +x_35 = lean_array_uset(x_24, x_28, x_34); +x_36 = l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(x_32); +x_37 = lean_nat_dec_le(x_36, x_25); +lean_dec(x_25); +lean_dec(x_36); +if (x_37 == 0) { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +lean_object* x_38; +x_38 = l_Std_HashMapImp_expand___at_Lean_Compiler_Simp_OccInfo_add___spec__5(x_32, x_35); +return x_38; } else { -uint8_t x_26; -lean_dec(x_16); -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_21); -if (x_26 == 0) -{ -return x_21; +lean_object* x_39; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 1, x_35); +return x_39; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_25); +x_40 = l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(x_2, x_3, x_29); +x_41 = lean_array_uset(x_24, x_28, x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_23); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_OccInfo_add(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +lean_object* x_3; +lean_inc(x_2); +lean_inc(x_1); +x_3 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +uint8_t x_4; lean_object* x_5; +x_4 = 0; +x_5 = l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3(x_1, x_2, x_4); +return x_5; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +x_7 = lean_unbox(x_6); +lean_dec(x_6); +if (x_7 == 0) { -return x_31; +uint8_t x_8; lean_object* x_9; +x_8 = 1; +x_9 = l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3(x_1, x_2, x_8); +return x_9; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_dec(x_2); +return x_1; } } } } -else +LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); +lean_object* x_3; +x_3 = l_Std_AssocList_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__2(x_1, x_2); lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +lean_dec(x_1); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_AssocList_contains___at_Lean_Compiler_Simp_OccInfo_add___spec__4(x_1, x_2); lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; -} +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, 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_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_2); lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) -{ -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +x_5 = l_Std_AssocList_replace___at_Lean_Compiler_Simp_OccInfo_add___spec__8(x_1, x_4, x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_26; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); lean_dec(x_3); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) -{ -return x_21; +x_5 = l_Std_HashMap_insert___at_Lean_Compiler_Simp_OccInfo_add___spec__3(x_1, x_2, x_4); +return x_5; } -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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; } +static lean_object* _init_l_Lean_Compiler_Simp_Config_smallThreshold___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_unsigned_to_nat(1u); +return x_1; } } -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +static lean_object* _init_l_Lean_Compiler_Simp_Context_config___default() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_1; +x_1 = lean_unsigned_to_nat(1u); +return x_1; } -else +} +static lean_object* _init_l_Lean_Compiler_Simp_State_occInfo___default() { +_start: { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_object* x_1; +x_1 = l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1; +return x_1; +} +} +static uint8_t _init_l_Lean_Compiler_Simp_State_simplified___default() { +_start: { -return x_31; +uint8_t x_1; +x_1 = 0; +return x_1; } -else +} +static lean_object* _init_l_Lean_Compiler_Simp_instInhabitedState___closed__1() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1; +x_2 = 0; +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); +return x_3; } } +static lean_object* _init_l_Lean_Compiler_Simp_instInhabitedState() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Compiler_Simp_instInhabitedState___closed__1; +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitValue(lean_object* x_1, 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_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; -} +uint8_t x_8; +x_8 = l_Lean_Expr_isApp(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +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; } else { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_Lean_Expr_getAppFn(x_1); +x_12 = l_Lean_Compiler_Simp_findLambda_x3f(x_11, x_4, x_5, x_6, x_7); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +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; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 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_12, 1); +x_22 = lean_ctor_get(x_12, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_13, 0); +lean_inc(x_23); +lean_dec(x_13); +x_24 = l_Lean_LocalDecl_value(x_23); +x_25 = l_Lean_Expr_isLambda(x_24); +lean_dec(x_24); +if (x_25 == 0) { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +lean_object* x_26; +lean_dec(x_23); +x_26 = lean_box(0); +lean_ctor_set(x_12, 0, x_26); +return x_12; } else { -uint8_t x_26; -lean_dec(x_16); -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_21); -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; uint8_t x_33; +lean_free_object(x_12); +x_27 = l_Lean_LocalDecl_userName(x_23); +lean_dec(x_23); +x_28 = lean_st_ref_get(x_6, x_21); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_st_ref_take(x_3, 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); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_31, 0); +x_35 = l_Lean_Compiler_Simp_OccInfo_add(x_34, x_27); +lean_ctor_set(x_31, 0, x_35); +x_36 = lean_st_ref_set(x_3, x_31, x_32); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -return x_21; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_box(0); +lean_ctor_set(x_36, 0, x_39); +return x_36; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +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_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); +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; +x_43 = lean_ctor_get(x_31, 0); +x_44 = lean_ctor_get_uint8(x_31, sizeof(void*)*1); +lean_inc(x_43); lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +x_45 = l_Lean_Compiler_Simp_OccInfo_add(x_43, x_27); +x_46 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*1, x_44); +x_47 = lean_st_ref_set(x_3, x_46, x_32); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_12, 1); +lean_inc(x_52); +lean_dec(x_12); +x_53 = lean_ctor_get(x_13, 0); +lean_inc(x_53); +lean_dec(x_13); +x_54 = l_Lean_LocalDecl_value(x_53); +x_55 = l_Lean_Expr_isLambda(x_54); +lean_dec(x_54); +if (x_55 == 0) { -return x_31; +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_52); +return x_57; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +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; 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; lean_object* x_73; +x_58 = l_Lean_LocalDecl_userName(x_53); +lean_dec(x_53); +x_59 = lean_st_ref_get(x_6, x_52); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = lean_st_ref_take(x_3, x_60); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +x_65 = lean_ctor_get_uint8(x_62, sizeof(void*)*1); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + x_66 = x_62; +} else { + lean_dec_ref(x_62); + x_66 = lean_box(0); } +x_67 = l_Lean_Compiler_Simp_OccInfo_add(x_64, x_58); +if (lean_is_scalar(x_66)) { + x_68 = lean_alloc_ctor(0, 1, 1); +} else { + x_68 = x_66; } +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set_uint8(x_68, sizeof(void*)*1, x_65); +x_69 = lean_st_ref_set(x_3, x_68, x_63); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_71 = x_69; +} else { + lean_dec_ref(x_69); + x_71 = lean_box(0); } +x_72 = lean_box(0); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = x_71; } -else -{ -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_70); +return x_73; } } -else -{ -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_10, lean_object* x_11) { +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitValue___boxed(lean_object* x_1, 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_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_internalize_visitValue(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +lean_dec(x_1); +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_internalize_visitLambda___closed__1() { +_start: { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _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; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLambda(lean_object* x_1, 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_26; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_8; 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +x_10 = lean_st_ref_get(x_4, x_9); +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_41 = lean_st_ref_get(x_6, x_12); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_st_ref_take(x_4, x_42); +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_is_exclusive(x_44); +if (x_46 == 0) { -return x_21; +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_47 = lean_ctor_get(x_44, 1); +lean_dec(x_47); +x_48 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +lean_ctor_set(x_44, 1, x_48); +x_49 = lean_st_ref_set(x_4, x_44, x_45); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_Lean_Compiler_visitLambdaCore(x_1, x_4, x_5, x_6, x_50); +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, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +lean_dec(x_52); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_54); +x_56 = l_Lean_Compiler_Simp_internalize_visitLet(x_55, x_54, x_2, x_3, x_4, x_5, x_6, x_53); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +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); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_59 = l_Lean_Compiler_mkLetUsingScope(x_57, x_4, x_5, x_6, x_58); +if (lean_obj_tag(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; lean_object* x_66; 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_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = l_Lean_Compiler_mkLambda(x_54, x_60, x_4, x_5, x_6, x_61); +lean_dec(x_5); +lean_dec(x_54); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_st_ref_get(x_6, x_64); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +x_67 = lean_st_ref_get(x_4, x_66); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_ctor_get(x_11, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_11, 1); +lean_inc(x_71); +lean_dec(x_11); +x_72 = !lean_is_exclusive(x_68); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_73 = lean_ctor_get(x_68, 1); +lean_dec(x_73); +x_74 = lean_ctor_get(x_68, 0); +lean_dec(x_74); +lean_ctor_set(x_68, 1, x_71); +lean_ctor_set(x_68, 0, x_70); +x_75 = lean_st_ref_get(x_6, x_69); +lean_dec(x_6); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_st_ref_set(x_4, x_68, x_76); +lean_dec(x_4); +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_77, 0); +lean_dec(x_79); +lean_ctor_set(x_77, 0, x_63); +return x_77; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_63); +lean_ctor_set(x_81, 1, x_80); +return x_81; } } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +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; +x_82 = lean_ctor_get(x_68, 2); +lean_inc(x_82); +lean_dec(x_68); +x_83 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_83, 0, x_70); +lean_ctor_set(x_83, 1, x_71); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_st_ref_get(x_6, x_69); +lean_dec(x_6); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = lean_st_ref_set(x_4, x_83, x_85); +lean_dec(x_4); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_88 = x_86; +} else { + lean_dec_ref(x_86); + x_88 = lean_box(0); +} +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_88; +} +lean_ctor_set(x_89, 0, x_63); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_object* x_90; lean_object* x_91; +lean_dec(x_54); +lean_dec(x_5); +x_90 = lean_ctor_get(x_59, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_59, 1); +lean_inc(x_91); +lean_dec(x_59); +x_13 = x_90; +x_14 = x_91; +goto block_40; +} +} +else { -return x_31; +lean_object* x_92; lean_object* x_93; +lean_dec(x_54); +lean_dec(x_5); +x_92 = lean_ctor_get(x_56, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_56, 1); +lean_inc(x_93); +lean_dec(x_56); +x_13 = x_92; +x_14 = x_93; +goto block_40; +} } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +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; +x_94 = lean_ctor_get(x_44, 0); +x_95 = lean_ctor_get(x_44, 2); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_44); +x_96 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_97 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); +x_98 = lean_st_ref_set(x_4, x_97, x_45); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = l_Lean_Compiler_visitLambdaCore(x_1, x_4, x_5, x_6, x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_101, 1); +lean_inc(x_104); +lean_dec(x_101); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_103); +x_105 = l_Lean_Compiler_Simp_internalize_visitLet(x_104, x_103, x_2, x_3, x_4, x_5, x_6, x_102); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_108 = l_Lean_Compiler_mkLetUsingScope(x_106, x_4, x_5, x_6, x_107); +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; 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; +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 = l_Lean_Compiler_mkLambda(x_103, x_109, x_4, x_5, x_6, x_110); +lean_dec(x_5); +lean_dec(x_103); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_st_ref_get(x_6, x_113); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_get(x_4, x_115); +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_11, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_11, 1); +lean_inc(x_120); +lean_dec(x_11); +x_121 = lean_ctor_get(x_117, 2); +lean_inc(x_121); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + 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, 3, 0); +} else { + x_123 = x_122; } +lean_ctor_set(x_123, 0, x_119); +lean_ctor_set(x_123, 1, x_120); +lean_ctor_set(x_123, 2, x_121); +x_124 = lean_st_ref_get(x_6, x_118); +lean_dec(x_6); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_126 = lean_st_ref_set(x_4, x_123, x_125); +lean_dec(x_4); +x_127 = lean_ctor_get(x_126, 1); +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); } +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_112); +lean_ctor_set(x_129, 1, x_127); +return x_129; } else { -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +lean_object* x_130; lean_object* x_131; +lean_dec(x_103); +lean_dec(x_5); +x_130 = lean_ctor_get(x_108, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_108, 1); +lean_inc(x_131); +lean_dec(x_108); +x_13 = x_130; +x_14 = x_131; +goto block_40; } } else { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; -} +lean_object* x_132; lean_object* x_133; +lean_dec(x_103); +lean_dec(x_5); +x_132 = lean_ctor_get(x_105, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_105, 1); +lean_inc(x_133); +lean_dec(x_105); +x_13 = x_132; +x_14 = x_133; +goto block_40; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) +block_40: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); +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_15 = lean_st_ref_get(x_6, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_4, 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); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) -{ -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; -} -else +x_20 = lean_ctor_get(x_11, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_dec(x_11); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -uint8_t x_26; -lean_dec(x_16); -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_21); -if (x_26 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_23 = lean_ctor_get(x_18, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_18, 0); +lean_dec(x_24); +lean_ctor_set(x_18, 1, x_21); +lean_ctor_set(x_18, 0, x_20); +x_25 = lean_st_ref_get(x_6, x_19); +lean_dec(x_6); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_set(x_4, x_18, x_26); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -return x_21; +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 0, x_13); +return x_27; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_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(1, 2, 0); +lean_ctor_set(x_31, 0, x_13); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); +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_32 = lean_ctor_get(x_18, 2); +lean_inc(x_32); +lean_dec(x_18); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_21); +lean_ctor_set(x_33, 2, x_32); +x_34 = lean_st_ref_get(x_6, x_19); +lean_dec(x_6); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_st_ref_set(x_4, x_33, x_35); +lean_dec(x_4); +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); +} +if (lean_is_scalar(x_38)) { + x_39 = lean_alloc_ctor(1, 2, 0); +} else { + x_39 = x_38; + lean_ctor_set_tag(x_39, 1); +} +lean_ctor_set(x_39, 0, x_13); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLet___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_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_7); +x_14 = l_Lean_Compiler_mkLetDecl(x_1, x_2, x_6, x_3, x_10, x_11, x_12, x_13); +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_array_push(x_4, x_15); +x_18 = l_Lean_Compiler_Simp_internalize_visitLet(x_5, x_17, x_8, x_9, x_10, x_11, x_12, x_16); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLet(lean_object* x_1, lean_object* x_2, 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) == 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; uint8_t x_19; lean_object* x_20; +x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); -lean_inc(x_8); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_12); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_dec(x_1); +x_14 = l_Lean_Compiler_mkFreshLetVarIdx(x_5, x_6, x_7, x_8); +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_expr_instantiate_rev(x_10, x_2); +lean_dec(x_10); +x_18 = lean_expr_instantiate_rev(x_11, x_2); +lean_dec(x_11); +x_19 = l_Lean_Expr_isLambda(x_18); +if (lean_obj_tag(x_9) == 2) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_9, 0); +lean_inc(x_35); +lean_dec(x_9); +x_36 = l_Lean_Name_num___override(x_35, x_15); +x_20 = x_36; +goto block_34; +} +else +{ +lean_object* x_37; +x_37 = l_Lean_Name_num___override(x_9, x_15); +x_20 = x_37; +goto block_34; +} +block_34: +{ +if (x_19 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = l_Lean_Compiler_Simp_internalize_visitValue(x_18, x_3, x_4, x_5, x_6, x_7, x_16); +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 = l_Lean_Compiler_Simp_internalize_visitLet___lambda__1(x_20, x_17, x_13, x_2, x_12, x_18, x_22, x_3, x_4, x_5, x_6, x_7, x_23); +return x_24; +} +else +{ +lean_object* x_25; lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_25 = l_Lean_Compiler_Simp_internalize_visitLambda(x_18, x_3, x_4, x_5, x_6, x_7, x_16); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_box(0); +x_29 = l_Lean_Compiler_Simp_internalize_visitLet___lambda__1(x_20, x_17, x_13, x_2, x_12, x_26, x_28, x_3, x_4, x_5, x_6, x_7, x_27); +return x_29; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_30; +lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_12); lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_25); +if (x_30 == 0) { -return x_31; +return x_25; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_25, 0); +x_32 = lean_ctor_get(x_25, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_25); +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; +} } } } } else { +lean_object* x_38; lean_object* x_39; +x_38 = lean_expr_instantiate_rev(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_38); +x_39 = l_Lean_Compiler_isCasesApp_x3f(x_38, x_6, x_7, x_8); +if (lean_obj_tag(x_39) == 0) +{ lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = l_Lean_Compiler_Simp_internalize_visitValue(x_38, x_3, x_4, x_5, x_6, x_7, x_41); 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_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +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_42, 0, x_38); +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); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_38); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } else { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_48 = lean_ctor_get(x_40, 0); +lean_inc(x_48); +lean_dec(x_40); +x_49 = l_Lean_Compiler_Simp_internalize_visitCases(x_48, x_38, x_3, x_4, x_5, x_6, x_7, x_47); +return x_49; +} +} +else +{ +uint8_t x_50; +lean_dec(x_38); 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_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +x_50 = !lean_is_exclusive(x_39); +if (x_50 == 0) +{ +return x_39; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_39, 0); +x_52 = lean_ctor_get(x_39, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_39); +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; +} +} } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___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) { _start: { uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); +x_12 = lean_nat_dec_le(x_3, x_2); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); +x_14 = lean_nat_dec_eq(x_1, 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_6); x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); +x_16 = lean_nat_sub(x_1, x_15); +lean_dec(x_1); +x_17 = lean_array_get_size(x_5); +x_18 = lean_nat_dec_lt(x_2, x_17); lean_dec(x_17); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); +lean_object* x_19; +x_19 = lean_nat_add(x_2, x_4); +lean_dec(x_2); +x_1 = x_16; +x_2 = x_19; +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_array_fget(x_5, x_2); +x_22 = lean_box(0); +x_23 = lean_array_fset(x_5, x_2, x_22); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +lean_inc(x_6); +x_24 = l_Lean_Compiler_Simp_internalize_visitLambda(x_21, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_24) == 0) { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* 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 = lean_array_fset(x_23, x_2, x_25); +x_28 = lean_nat_add(x_2, x_4); +lean_dec(x_2); +x_1 = x_16; +x_2 = x_28; +x_5 = x_27; +x_11 = x_26; goto _start; } else { -uint8_t x_26; +uint8_t x_30; +lean_dec(x_23); lean_dec(x_16); 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_21); -if (x_26 == 0) +lean_dec(x_6); +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_24); +if (x_30 == 0) { -return x_21; +return x_24; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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; -} -} -} -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_24, 0); +x_32 = lean_ctor_get(x_24, 1); lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; -} -else -{ -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) -{ -return x_31; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_inc(x_31); +lean_dec(x_24); +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; } } } } else { -lean_object* x_40; +lean_object* x_34; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); +lean_dec(x_6); lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +lean_dec(x_1); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_5); +lean_ctor_set(x_34, 1, x_11); +return x_34; } } else { -lean_object* x_41; +lean_object* x_35; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); +lean_dec(x_6); lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +lean_dec(x_1); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_5); +lean_ctor_set(x_35, 1, x_11); +return x_35; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Compiler_Simp_internalize_visitCases___closed__1() { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitCases(lean_object* x_1, lean_object* x_2, 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_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); +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_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_9); +x_11 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_10); +x_12 = lean_mk_array(x_10, x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_10, x_13); +lean_dec(x_10); +lean_inc(x_2); +x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_12, x_14); +x_16 = lean_ctor_get(x_1, 3); +lean_inc(x_16); +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_16, 2); +lean_inc(x_19); +lean_dec(x_16); +lean_inc(x_17); +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___spec__1(x_17, x_18, x_17, x_19, x_15, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_19); lean_dec(x_17); -if (x_18 == 0) +if (lean_obj_tag(x_20) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { 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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; -} -else -{ -uint8_t x_26; -lean_dec(x_16); -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_21); -if (x_26 == 0) -{ -return x_21; +x_22 = lean_ctor_get(x_20, 0); +x_23 = l_Lean_Expr_getAppFn(x_2); +lean_dec(x_2); +x_24 = l_Lean_mkAppN(x_23, x_22); +lean_ctor_set(x_20, 0, x_24); +return x_20; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); +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_20, 0); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_20); +x_27 = l_Lean_Expr_getAppFn(x_2); +lean_dec(x_2); +x_28 = l_Lean_mkAppN(x_27, x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); return x_29; } } -} -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; -} else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) { -return x_31; +return x_20; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_20); +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; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize_visitLet___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_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +x_15 = l_Lean_Compiler_Simp_internalize_visitLet___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); +return x_15; } } -else +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___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: { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_internalize_visitCases___spec__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); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_internalize(lean_object* x_1, 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_Compiler_Simp_internalize_visitLambda(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_8; +} } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___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_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 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_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_st_ref_take(x_1, 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_is_exclusive(x_9); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); +uint8_t x_12; lean_object* x_13; uint8_t x_14; +x_12 = 1; +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_12); +x_13 = lean_st_ref_set(x_1, x_9, x_10); +x_14 = !lean_is_exclusive(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_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) -{ -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_13, 1); +lean_inc(x_17); +lean_dec(x_13); +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 { -uint8_t x_26; -lean_dec(x_16); -lean_dec(x_10); +lean_object* x_20; uint8_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; +x_20 = lean_ctor_get(x_9, 0); +lean_inc(x_20); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) -{ -return x_21; +x_21 = 1; +x_22 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set_uint8(x_22, sizeof(void*)*1, x_21); +x_23 = lean_st_ref_set(x_1, x_22, x_10); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_25 = x_23; +} else { + lean_dec_ref(x_23); + x_25 = lean_box(0); } -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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; +x_26 = lean_box(0); +if (lean_is_scalar(x_25)) { + x_27 = lean_alloc_ctor(0, 2, 0); +} else { + x_27 = x_25; } +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; } } -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified(lean_object* x_1) { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_markSimplified___rarg___boxed), 5, 0); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___rarg___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_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_6; +x_6 = l_Lean_Compiler_Simp_markSimplified___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) -{ -return x_31; +lean_dec(x_2); +lean_dec(x_1); +return x_6; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___boxed(lean_object* x_1) { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_object* x_2; +x_2 = l_Lean_Compiler_Simp_markSimplified(x_1); +lean_dec(x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findCtor(lean_object* x_1, 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 = 1; +x_9 = l_Lean_Compiler_Simp_findExpr(x_1, x_8, x_4, x_5, x_6, x_7); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_findCtor___boxed(lean_object* x_1, 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_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_findCtor(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); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +return x_8; } } -else +static lean_object* _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__1() { +_start: { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Init.Util", 9); +return x_1; } } +static lean_object* _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("getElem!", 8); +return x_1; +} } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__3() { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("index out of bounds", 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__4() { +_start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 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_Compiler_Simp_simpProj_x3f___closed__1; +x_2 = l_Lean_Compiler_Simp_simpProj_x3f___closed__2; +x_3 = lean_unsigned_to_nat(73u); +x_4 = lean_unsigned_to_nat(36u); +x_5 = l_Lean_Compiler_Simp_simpProj_x3f___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_Compiler_Simp_simpProj_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_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +if (lean_obj_tag(x_1) == 11) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); +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_1, 1); lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +lean_dec(x_1); +x_10 = 1; +x_11 = l_Lean_Compiler_Simp_findExpr(x_9, x_10, x_4, x_5, x_6, x_7); +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_st_ref_get(x_6, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Expr_constructorApp_x3f(x_18, x_12); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +lean_dec(x_8); +x_20 = lean_box(0); +lean_ctor_set(x_14, 0, x_20); +return x_14; } else { -uint8_t x_26; -lean_dec(x_16); -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_21); +uint8_t x_21; +lean_free_object(x_14); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +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_19, 0); +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 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_17); +x_26 = !lean_is_exclusive(x_25); if (x_26 == 0) { -return x_21; +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_25, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_23, 3); +lean_inc(x_28); +lean_dec(x_23); +x_29 = lean_nat_add(x_28, x_8); +lean_dec(x_8); +lean_dec(x_28); +x_30 = lean_array_get_size(x_24); +x_31 = lean_nat_dec_lt(x_29, x_30); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); +lean_dec(x_24); +x_32 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_33 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_32); +lean_ctor_set(x_19, 0, x_33); +lean_ctor_set(x_25, 0, x_19); +return x_25; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_object* x_34; +x_34 = lean_array_fget(x_24, x_29); +lean_dec(x_29); +lean_dec(x_24); +lean_ctor_set(x_19, 0, x_34); +lean_ctor_set(x_25, 0, x_19); +return x_25; } } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 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_25, 1); +lean_inc(x_35); +lean_dec(x_25); +x_36 = lean_ctor_get(x_23, 3); +lean_inc(x_36); +lean_dec(x_23); +x_37 = lean_nat_add(x_36, x_8); +lean_dec(x_8); +lean_dec(x_36); +x_38 = lean_array_get_size(x_24); +x_39 = lean_nat_dec_lt(x_37, x_38); +lean_dec(x_38); +if (x_39 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_37); +lean_dec(x_24); +x_40 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_41 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_40); +lean_ctor_set(x_19, 0, x_41); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_19); +lean_ctor_set(x_42, 1, x_35); +return x_42; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_43; lean_object* x_44; +x_43 = lean_array_fget(x_24, x_37); +lean_dec(x_37); +lean_dec(x_24); +lean_ctor_set(x_19, 0, x_43); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_19); +lean_ctor_set(x_44, 1, x_35); +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; uint8_t x_54; +x_45 = lean_ctor_get(x_19, 0); +lean_inc(x_45); +lean_dec(x_19); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_17); +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_ctor_get(x_46, 3); +lean_inc(x_51); +lean_dec(x_46); +x_52 = lean_nat_add(x_51, x_8); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_dec(x_51); +x_53 = lean_array_get_size(x_47); +x_54 = lean_nat_dec_lt(x_52, x_53); +lean_dec(x_53); +if (x_54 == 0) { -return x_31; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_52); +lean_dec(x_47); +x_55 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_56 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_55); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_56); +if (lean_is_scalar(x_50)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_50; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_49); +return x_58; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_array_fget(x_47, x_52); +lean_dec(x_52); +lean_dec(x_47); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_59); +if (lean_is_scalar(x_50)) { + x_61 = lean_alloc_ctor(0, 2, 0); +} else { + x_61 = x_50; +} +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_49); +return x_61; } } } } else { -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_14, 0); +x_63 = lean_ctor_get(x_14, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_14); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_Lean_Expr_constructorApp_x3f(x_64, x_12); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; -} +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_63); +return x_67; } else { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); +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; +x_68 = lean_ctor_get(x_65, 0); +lean_inc(x_68); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + x_69 = x_65; +} else { + lean_dec_ref(x_65); + x_69 = lean_box(0); +} +x_70 = lean_ctor_get(x_68, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_68, 1); +lean_inc(x_71); +lean_dec(x_68); +x_72 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_63); +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_74 = x_72; +} else { + lean_dec_ref(x_72); + x_74 = lean_box(0); +} +x_75 = lean_ctor_get(x_70, 3); +lean_inc(x_75); +lean_dec(x_70); +x_76 = lean_nat_add(x_75, x_8); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; +lean_dec(x_75); +x_77 = lean_array_get_size(x_71); +x_78 = lean_nat_dec_lt(x_76, x_77); +lean_dec(x_77); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_76); +lean_dec(x_71); +x_79 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_80 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_79); +if (lean_is_scalar(x_69)) { + x_81 = lean_alloc_ctor(1, 1, 0); +} else { + x_81 = x_69; } +lean_ctor_set(x_81, 0, x_80); +if (lean_is_scalar(x_74)) { + x_82 = lean_alloc_ctor(0, 2, 0); +} else { + x_82 = x_74; } +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_73); +return x_82; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: +else { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_array_fget(x_71, x_76); +lean_dec(x_76); +lean_dec(x_71); +if (lean_is_scalar(x_69)) { + x_84 = lean_alloc_ctor(1, 1, 0); +} else { + x_84 = x_69; +} +lean_ctor_set(x_84, 0, x_83); +if (lean_is_scalar(x_74)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_74; +} +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_73); +return x_85; +} +} +} +} +else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) +lean_object* x_86; lean_object* x_87; +lean_dec(x_1); +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_7); +return x_87; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpProj_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) { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_simpProj_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1() { +_start: { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _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; } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpAppApp_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: { -uint8_t x_26; -lean_dec(x_16); -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_21); -if (x_26 == 0) +uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_Lean_Expr_isApp(x_1); +if (x_8 == 0) { -return x_21; +lean_object* x_139; +x_139 = lean_box(0); +x_9 = x_139; +x_10 = x_7; +goto block_138; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_object* x_140; +x_140 = l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1; +x_9 = x_140; +x_10 = x_7; +goto block_138; } +block_138: +{ +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_11; lean_object* x_12; +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 +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_9, 0); +lean_dec(x_14); +x_15 = l_Lean_Expr_getAppFn(x_1); +x_16 = l_Lean_Expr_isFVar(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_15); +lean_free_object(x_9); +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_10); +return x_18; } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +uint8_t x_19; lean_object* x_20; uint8_t x_21; +x_19 = 1; +x_20 = l_Lean_Compiler_Simp_findExpr(x_15, x_19, x_4, x_5, x_6, x_10); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = l_Lean_Expr_isApp(x_22); +if (x_24 == 0) +{ +uint8_t x_25; +x_25 = l_Lean_Expr_isConst(x_22); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_22); +lean_free_object(x_9); +lean_dec(x_1); +x_26 = lean_box(0); +lean_ctor_set(x_20, 0, x_26); +return x_20; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_object* x_27; uint8_t x_28; +lean_free_object(x_20); +x_27 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_23); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -return x_31; +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 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_unsigned_to_nat(0u); +x_31 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_30); +x_32 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_31); +x_33 = lean_mk_array(x_31, x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_31, x_34); +lean_dec(x_31); +x_36 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_33, x_35); +x_37 = l_Lean_mkAppN(x_22, x_36); +lean_ctor_set(x_9, 0, x_37); +lean_ctor_set(x_27, 0, x_9); +return x_27; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); +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_38 = lean_ctor_get(x_27, 1); lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +lean_dec(x_27); +x_39 = lean_unsigned_to_nat(0u); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_39); +x_41 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_40); +x_42 = lean_mk_array(x_40, x_41); +x_43 = lean_unsigned_to_nat(1u); +x_44 = lean_nat_sub(x_40, x_43); +lean_dec(x_40); +x_45 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_42, x_44); +x_46 = l_Lean_mkAppN(x_22, x_45); +lean_ctor_set(x_9, 0, x_46); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_9); +lean_ctor_set(x_47, 1, x_38); +return x_47; } } } +else +{ +lean_object* x_48; uint8_t x_49; +lean_free_object(x_20); +x_48 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_23); +x_49 = !lean_is_exclusive(x_48); +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_ctor_get(x_48, 0); +lean_dec(x_50); +x_51 = lean_unsigned_to_nat(0u); +x_52 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_51); +x_53 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_52); +x_54 = lean_mk_array(x_52, x_53); +x_55 = lean_unsigned_to_nat(1u); +x_56 = lean_nat_sub(x_52, x_55); +lean_dec(x_52); +x_57 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_54, x_56); +x_58 = l_Lean_mkAppN(x_22, x_57); +lean_ctor_set(x_9, 0, x_58); +lean_ctor_set(x_48, 0, x_9); +return x_48; } else { -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +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_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +lean_dec(x_48); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_60); +x_62 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_61); +x_63 = lean_mk_array(x_61, x_62); +x_64 = lean_unsigned_to_nat(1u); +x_65 = lean_nat_sub(x_61, x_64); +lean_dec(x_61); +x_66 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_63, x_65); +x_67 = l_Lean_mkAppN(x_22, x_66); +lean_ctor_set(x_9, 0, x_67); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_9); +lean_ctor_set(x_68, 1, x_59); +return x_68; +} } } else { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, 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_6); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = lean_array_get_size(x_1); -x_18 = lean_nat_dec_lt(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_69 = lean_ctor_get(x_20, 0); +x_70 = lean_ctor_get(x_20, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_20); +x_71 = l_Lean_Expr_isApp(x_69); +if (x_71 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_20 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_19); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_20, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_21) == 0) +uint8_t x_72; +x_72 = l_Lean_Expr_isConst(x_69); +if (x_72 == 0) { -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 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_24 = lean_box(0); -x_2 = x_16; -x_3 = x_23; -x_6 = x_24; -x_11 = x_22; -goto _start; +lean_object* x_73; lean_object* x_74; +lean_dec(x_69); +lean_free_object(x_9); +lean_dec(x_1); +x_73 = lean_box(0); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_70); +return x_74; } else { -uint8_t x_26; -lean_dec(x_16); -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_21); -if (x_26 == 0) -{ -return x_21; +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; +x_75 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_70); +x_76 = lean_ctor_get(x_75, 1); +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 = lean_unsigned_to_nat(0u); +x_79 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_78); +x_80 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_79); +x_81 = lean_mk_array(x_79, x_80); +x_82 = lean_unsigned_to_nat(1u); +x_83 = lean_nat_sub(x_79, x_82); +lean_dec(x_79); +x_84 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_81, x_83); +x_85 = l_Lean_mkAppN(x_69, x_84); +lean_ctor_set(x_9, 0, x_85); +if (lean_is_scalar(x_77)) { + x_86 = lean_alloc_ctor(0, 2, 0); +} else { + x_86 = x_77; +} +lean_ctor_set(x_86, 0, x_9); +lean_ctor_set(x_86, 1, x_76); +return x_86; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -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_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_87 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_70); +x_88 = lean_ctor_get(x_87, 1); +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 = lean_unsigned_to_nat(0u); +x_91 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_90); +x_92 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_91); +x_93 = lean_mk_array(x_91, x_92); +x_94 = lean_unsigned_to_nat(1u); +x_95 = lean_nat_sub(x_91, x_94); +lean_dec(x_91); +x_96 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_93, x_95); +x_97 = l_Lean_mkAppN(x_69, x_96); +lean_ctor_set(x_9, 0, x_97); +if (lean_is_scalar(x_89)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_89; +} +lean_ctor_set(x_98, 0, x_9); +lean_ctor_set(x_98, 1, x_88); +return x_98; +} } } } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_array_fget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_31 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_30, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_31) == 0) +lean_object* x_99; uint8_t x_100; +lean_dec(x_9); +x_99 = l_Lean_Expr_getAppFn(x_1); +x_100 = l_Lean_Expr_isFVar(x_99); +if (x_100 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_34 = lean_box(0); -x_2 = x_16; -x_3 = x_33; -x_6 = x_34; -x_11 = x_32; -goto _start; +lean_object* x_101; lean_object* x_102; +lean_dec(x_99); +lean_dec(x_1); +x_101 = lean_box(0); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_10); +return x_102; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = 1; +x_104 = l_Lean_Compiler_Simp_findExpr(x_99, x_103, x_4, x_5, x_6, x_10); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; +} else { + lean_dec_ref(x_104); + x_107 = lean_box(0); +} +x_108 = l_Lean_Expr_isApp(x_105); +if (x_108 == 0) { -return x_31; +uint8_t x_109; +x_109 = l_Lean_Expr_isConst(x_105); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_105); +lean_dec(x_1); +x_110 = lean_box(0); +if (lean_is_scalar(x_107)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_107; +} +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_106); +return x_111; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_31); -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; +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_dec(x_107); +x_112 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_106); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); } +x_115 = lean_unsigned_to_nat(0u); +x_116 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_115); +x_117 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_116); +x_118 = lean_mk_array(x_116, x_117); +x_119 = lean_unsigned_to_nat(1u); +x_120 = lean_nat_sub(x_116, x_119); +lean_dec(x_116); +x_121 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_118, x_120); +x_122 = l_Lean_mkAppN(x_105, x_121); +x_123 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_123, 0, x_122); +if (lean_is_scalar(x_114)) { + x_124 = lean_alloc_ctor(0, 2, 0); +} else { + x_124 = x_114; } +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_113); +return x_124; } } else { -lean_object* x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_6); -lean_ctor_set(x_40, 1, x_11); -return x_40; +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_dec(x_107); +x_125 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_106); +x_126 = lean_ctor_get(x_125, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_127 = x_125; +} else { + lean_dec_ref(x_125); + x_127 = lean_box(0); +} +x_128 = lean_unsigned_to_nat(0u); +x_129 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_128); +x_130 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_129); +x_131 = lean_mk_array(x_129, x_130); +x_132 = lean_unsigned_to_nat(1u); +x_133 = lean_nat_sub(x_129, x_132); +lean_dec(x_129); +x_134 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_131, x_133); +x_135 = l_Lean_mkAppN(x_105, x_134); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_135); +if (lean_is_scalar(x_127)) { + x_137 = lean_alloc_ctor(0, 2, 0); +} else { + x_137 = x_127; } +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_126); +return x_137; } -else +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpAppApp_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) { +_start: { -lean_object* x_41; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_simpAppApp_x3f(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); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_6); -lean_ctor_set(x_41, 1, x_11); -return x_41; -} +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_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, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInlineLocal(lean_object* x_1, 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_inc(x_3); -x_8 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_3, x_1, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_LocalDecl_userName(x_1); +x_16 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__1(x_14, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_8, 0); -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_3); -return x_8; +lean_object* x_17; lean_object* x_18; +lean_free_object(x_10); +x_17 = l_Lean_LocalDecl_value(x_1); +x_18 = l_Lean_Compiler_lcnfSizeLe(x_17, x_2, x_5, x_6, x_13); +return x_18; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_21 = 1; +x_22 = lean_box(x_21); +lean_ctor_set(x_10, 0, x_22); +return x_10; +} +else +{ +lean_object* x_23; lean_object* x_24; +lean_free_object(x_10); +x_23 = l_Lean_LocalDecl_value(x_1); +x_24 = l_Lean_Compiler_lcnfSizeLe(x_23, x_2, x_5, x_6, x_13); +return x_24; +} } } else { -uint8_t x_13; -lean_dec(x_3); -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) +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_10, 0); +x_26 = lean_ctor_get(x_10, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_LocalDecl_userName(x_1); +x_29 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_Simp_OccInfo_add___spec__1(x_27, x_28); +if (lean_obj_tag(x_29) == 0) { -return x_8; +lean_object* x_30; lean_object* x_31; +x_30 = l_Lean_LocalDecl_value(x_1); +x_31 = l_Lean_Compiler_lcnfSizeLe(x_30, x_2, x_5, x_6, x_26); +return x_31; } 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_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_unbox(x_32); +lean_dec(x_32); +if (x_33 == 0) +{ +uint8_t x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_34 = 1; +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_26); +return x_36; } +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = l_Lean_LocalDecl_value(x_1); +x_38 = l_Lean_Compiler_lcnfSizeLe(x_37, x_2, x_5, x_6, x_26); +return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_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) { -_start: -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_collectInlineStats_go___lambda__1___boxed), 7, 1); -lean_closure_set(x_7, 0, x_2); -x_8 = l_Lean_Compiler_visitLetImp(x_1, x_7, x_3, x_4, x_5, x_6); -return x_8; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_go___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_shouldInlineLocal___boxed(lean_object* x_1, 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_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_shouldInlineLocal(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_8; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_go___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_collectInlineStats_go), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_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: +lean_object* x_10; +x_10 = l_Lean_Compiler_ensureUniqueLetVarNames(x_1, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -switch (lean_obj_tag(x_1)) { -case 0: +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_7; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_7 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_9); +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_10, 0); +x_13 = 1; +x_14 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_12); +lean_ctor_set_uint8(x_14, sizeof(void*)*2, x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_10, 0, x_15); 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; 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_7, 1); -lean_inc(x_11); -lean_dec(x_7); -x_12 = lean_ctor_get(x_8, 0); -lean_inc(x_12); -lean_dec(x_8); -x_13 = lean_unsigned_to_nat(0u); -x_14 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_13); -x_15 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_14); -x_16 = lean_mk_array(x_14, x_15); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_14, x_17); -lean_dec(x_14); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_16, x_18); -x_20 = lean_ctor_get(x_12, 3); -lean_inc(x_20); -lean_dec(x_12); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_20, 2); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_box(0); -lean_inc(x_21); -x_25 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1(x_19, x_21, x_22, x_21, x_23, x_24, x_2, x_3, x_4, x_5, x_11); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_19); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) -{ -lean_object* x_27; -x_27 = lean_ctor_get(x_25, 0); -lean_dec(x_27); -lean_ctor_set(x_25, 0, x_24); -return x_25; -} -else -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_24); -lean_ctor_set(x_29, 1, x_28); -return x_29; +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_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_10); +x_18 = 1; +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_16); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, 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_17); +return x_21; } } else { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_25); -if (x_30 == 0) +uint8_t x_22; +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { -return x_25; +return x_10; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_25, 0); -x_32 = lean_ctor_get(x_25, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_25); -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_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_10); +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_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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: { -uint8_t x_34; -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_dec(x_3); -lean_dec(x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_10); lean_dec(x_1); -x_34 = !lean_is_exclusive(x_7); -if (x_34 == 0) +x_12 = l_Lean_LocalDecl_value(x_2); +lean_dec(x_2); +x_13 = l_Lean_Compiler_getLambdaArity(x_12); +x_14 = lean_nat_dec_lt(x_11, x_13); +lean_dec(x_11); +if (x_14 == 0) { -return x_7; +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__1(x_12, x_13, x_15, x_4, x_5, x_6, x_7, x_8, x_9); +return x_16; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_7, 0); -x_36 = lean_ctor_get(x_7, 1); -lean_inc(x_36); -lean_inc(x_35); +lean_object* x_17; lean_object* x_18; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); lean_dec(x_7); -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; -} +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_9); +return x_18; } } -case 1: -{ -lean_object* x_38; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_38 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; -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; -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_40); -return x_41; } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_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; -x_42 = lean_ctor_get(x_38, 1); -lean_inc(x_42); -lean_dec(x_38); -x_43 = lean_ctor_get(x_39, 0); -lean_inc(x_43); -lean_dec(x_39); -x_44 = lean_unsigned_to_nat(0u); -x_45 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_44); -x_46 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_45); -x_47 = lean_mk_array(x_45, x_46); -x_48 = lean_unsigned_to_nat(1u); -x_49 = lean_nat_sub(x_45, x_48); -lean_dec(x_45); -x_50 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_47, x_49); -x_51 = lean_ctor_get(x_43, 3); -lean_inc(x_51); -lean_dec(x_43); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_51, 2); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_box(0); -lean_inc(x_52); -x_56 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__2(x_50, x_52, x_53, x_52, x_54, x_55, x_2, x_3, x_4, x_5, x_42); -lean_dec(x_54); -lean_dec(x_52); -lean_dec(x_50); -if (lean_obj_tag(x_56) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_1, 3); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +lean_dec(x_1); +x_13 = l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(x_12, x_2, x_11); +lean_dec(x_12); +x_14 = l_Lean_Compiler_ensureUniqueLetVarNames(x_13, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_58; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -lean_ctor_set(x_56, 0, x_55); -return x_56; +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = 0; +x_18 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_18, 0, x_3); +lean_ctor_set(x_18, 1, x_16); +lean_ctor_set_uint8(x_18, sizeof(void*)*2, x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_14, 0, x_19); +return x_14; } else { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_55); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_14, 0); +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_14); +x_22 = 0; +x_23 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_23, 0, x_3); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set_uint8(x_23, sizeof(void*)*2, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, 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); +return x_25; } } else { -uint8_t x_61; -x_61 = !lean_is_exclusive(x_56); -if (x_61 == 0) +uint8_t x_26; +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_14); +if (x_26 == 0) { -return x_56; +return x_14; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_56, 0); -x_63 = lean_ctor_get(x_56, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_56); -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; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_14, 0); +x_28 = lean_ctor_get(x_14, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_14); +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; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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: { -uint8_t x_65; -lean_dec(x_5); +lean_object* x_11; lean_dec(x_4); +lean_inc(x_9); +lean_inc(x_8); +x_11 = l_Lean_Compiler_getStage1Decl_x3f(x_1, x_8, x_9, x_10); +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_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_65 = !lean_is_exclusive(x_38); -if (x_65 == 0) +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) { -return x_38; +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_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_38, 0); -x_67 = lean_ctor_get(x_38, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_38); -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; +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; } } -} -case 2: -{ -lean_object* x_69; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_69 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_69) == 0) -{ -lean_object* x_70; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -if (lean_obj_tag(x_70) == 0) -{ -lean_object* x_71; lean_object* x_72; -x_71 = lean_ctor_get(x_69, 1); -lean_inc(x_71); -lean_dec(x_69); -x_72 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, 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; 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_73 = lean_ctor_get(x_69, 1); -lean_inc(x_73); -lean_dec(x_69); -x_74 = lean_ctor_get(x_70, 0); -lean_inc(x_74); -lean_dec(x_70); -x_75 = lean_unsigned_to_nat(0u); -x_76 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_75); -x_77 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_76); -x_78 = lean_mk_array(x_76, x_77); -x_79 = lean_unsigned_to_nat(1u); -x_80 = lean_nat_sub(x_76, x_79); -lean_dec(x_76); -x_81 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_78, x_80); -x_82 = lean_ctor_get(x_74, 3); -lean_inc(x_82); -lean_dec(x_74); -x_83 = lean_ctor_get(x_82, 1); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_82, 2); -lean_inc(x_85); -lean_dec(x_82); -x_86 = lean_box(0); -lean_inc(x_83); -x_87 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__3(x_81, x_83, x_84, x_83, x_85, x_86, x_2, x_3, x_4, x_5, x_73); -lean_dec(x_85); -lean_dec(x_83); -lean_dec(x_81); -if (lean_obj_tag(x_87) == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_87); -if (x_88 == 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; uint8_t x_26; +x_20 = lean_ctor_get(x_11, 1); +x_21 = lean_ctor_get(x_11, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_12, 0); +lean_inc(x_22); +lean_dec(x_12); +x_23 = lean_unsigned_to_nat(0u); +x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_23); +lean_dec(x_2); +x_25 = l_Lean_Compiler_Decl_getArity(x_22); +x_26 = lean_nat_dec_lt(x_24, x_25); +lean_dec(x_24); +if (x_26 == 0) { -lean_object* x_89; -x_89 = lean_ctor_get(x_87, 0); -lean_dec(x_89); -lean_ctor_set(x_87, 0, x_86); -return x_87; +lean_object* x_27; lean_object* x_28; +lean_free_object(x_11); +x_27 = lean_box(0); +x_28 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3(x_22, x_3, x_25, x_27, x_5, x_6, x_7, x_8, x_9, x_20); +lean_dec(x_3); +return x_28; } else { -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_87, 1); -lean_inc(x_90); -lean_dec(x_87); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_86); -lean_ctor_set(x_91, 1, x_90); -return x_91; +lean_object* x_29; +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_29 = lean_box(0); +lean_ctor_set(x_11, 0, x_29); +return x_11; } } else { -uint8_t x_92; -x_92 = !lean_is_exclusive(x_87); -if (x_92 == 0) +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_11, 1); +lean_inc(x_30); +lean_dec(x_11); +x_31 = lean_ctor_get(x_12, 0); +lean_inc(x_31); +lean_dec(x_12); +x_32 = lean_unsigned_to_nat(0u); +x_33 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_32); +lean_dec(x_2); +x_34 = l_Lean_Compiler_Decl_getArity(x_31); +x_35 = lean_nat_dec_lt(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) { -return x_87; +lean_object* x_36; lean_object* x_37; +x_36 = lean_box(0); +x_37 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3(x_31, x_3, x_34, x_36, x_5, x_6, x_7, x_8, x_9, x_30); +lean_dec(x_3); +return x_37; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_87, 0); -x_94 = lean_ctor_get(x_87, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_87); -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_object* x_38; lean_object* x_39; +lean_dec(x_34); +lean_dec(x_31); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +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_30); +return x_39; } } } } else { -uint8_t x_96; -lean_dec(x_5); -lean_dec(x_4); +uint8_t x_40; +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_96 = !lean_is_exclusive(x_69); -if (x_96 == 0) +x_40 = !lean_is_exclusive(x_11); +if (x_40 == 0) { -return x_69; +return x_11; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -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_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_41; lean_object* x_42; lean_object* x_43; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } -case 3: -{ -lean_object* x_100; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_100 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_100) == 0) -{ -lean_object* x_101; -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_100, 1); -lean_inc(x_102); -lean_dec(x_100); -x_103 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_102); -return x_103; } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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_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; -x_104 = lean_ctor_get(x_100, 1); -lean_inc(x_104); -lean_dec(x_100); -x_105 = lean_ctor_get(x_101, 0); -lean_inc(x_105); -lean_dec(x_101); -x_106 = lean_unsigned_to_nat(0u); -x_107 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_106); -x_108 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_107); -x_109 = lean_mk_array(x_107, x_108); -x_110 = lean_unsigned_to_nat(1u); -x_111 = lean_nat_sub(x_107, x_110); -lean_dec(x_107); -x_112 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_109, x_111); -x_113 = lean_ctor_get(x_105, 3); -lean_inc(x_113); -lean_dec(x_105); -x_114 = lean_ctor_get(x_113, 1); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_113, 2); -lean_inc(x_116); -lean_dec(x_113); -x_117 = lean_box(0); -lean_inc(x_114); -x_118 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__4(x_112, x_114, x_115, x_114, x_116, x_117, x_2, x_3, x_4, x_5, x_104); -lean_dec(x_116); -lean_dec(x_114); -lean_dec(x_112); -if (lean_obj_tag(x_118) == 0) +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = l_Lean_Expr_getAppFn(x_1); +x_9 = 1; +lean_inc(x_8); +x_10 = l_Lean_Compiler_Simp_findExpr(x_8, x_9, x_4, x_5, x_6, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 4) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_dec(x_8); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_st_ref_get(x_6, x_12); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -uint8_t x_119; -x_119 = !lean_is_exclusive(x_118); -if (x_119 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 0; +lean_inc(x_13); +x_21 = l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(x_19, x_20, x_13); +if (x_21 == 0) { -lean_object* x_120; -x_120 = lean_ctor_get(x_118, 0); -lean_dec(x_120); -lean_ctor_set(x_118, 0, x_117); -return x_118; +lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +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_22 = lean_box(0); +lean_ctor_set(x_15, 0, x_22); +return x_15; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_118, 1); -lean_inc(x_121); -lean_dec(x_118); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_121); -return x_122; +lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_23 = lean_box(0); +x_24 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__4(x_13, x_1, x_14, x_23, x_2, x_3, x_4, x_5, x_6, x_18); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_24; } } else { -uint8_t x_123; -x_123 = !lean_is_exclusive(x_118); -if (x_123 == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; +x_25 = lean_ctor_get(x_15, 0); +x_26 = lean_ctor_get(x_15, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_15); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = 0; +lean_inc(x_13); +x_29 = l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(x_27, x_28, x_13); +if (x_29 == 0) { -return x_118; +lean_object* x_30; lean_object* x_31; +lean_dec(x_14); +lean_dec(x_13); +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_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_26); +return x_31; } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_124 = lean_ctor_get(x_118, 0); -x_125 = lean_ctor_get(x_118, 1); -lean_inc(x_125); -lean_inc(x_124); -lean_dec(x_118); -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; -} +lean_object* x_32; lean_object* x_33; +x_32 = lean_box(0); +x_33 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__4(x_13, x_1, x_14, x_32, x_2, x_3, x_4, x_5, x_6, x_26); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_33; } } } else { -uint8_t x_127; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_11); +x_34 = lean_ctor_get(x_10, 1); +lean_inc(x_34); +lean_dec(x_10); +x_35 = l_Lean_Compiler_Simp_findLambda_x3f(x_8, x_4, x_5, x_6, x_34); +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_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_127 = !lean_is_exclusive(x_100); -if (x_127 == 0) +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) { -return x_100; +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_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_100, 0); -x_129 = lean_ctor_get(x_100, 1); -lean_inc(x_129); -lean_inc(x_128); -lean_dec(x_100); -x_130 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -return x_130; -} +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; } } -case 4: +else { -lean_object* x_131; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +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); +lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_131 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_131) == 0) +lean_inc(x_2); +x_45 = l_Lean_Compiler_Simp_shouldInlineLocal(x_44, x_2, x_3, x_4, x_5, x_6, x_43); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_132; -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -if (lean_obj_tag(x_132) == 0) +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -lean_object* x_133; lean_object* x_134; -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_133); -return x_134; +uint8_t x_48; +lean_dec(x_44); +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_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_45, 0); +lean_dec(x_49); +x_50 = lean_box(0); +lean_ctor_set(x_45, 0, x_50); +return x_45; } 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; -x_135 = lean_ctor_get(x_131, 1); -lean_inc(x_135); -lean_dec(x_131); -x_136 = lean_ctor_get(x_132, 0); -lean_inc(x_136); -lean_dec(x_132); -x_137 = lean_unsigned_to_nat(0u); -x_138 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_137); -x_139 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_138); -x_140 = lean_mk_array(x_138, x_139); -x_141 = lean_unsigned_to_nat(1u); -x_142 = lean_nat_sub(x_138, x_141); -lean_dec(x_138); -x_143 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_140, x_142); -x_144 = lean_ctor_get(x_136, 3); -lean_inc(x_144); -lean_dec(x_136); -x_145 = lean_ctor_get(x_144, 1); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_144, 2); -lean_inc(x_147); -lean_dec(x_144); -x_148 = lean_box(0); -lean_inc(x_145); -x_149 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__5(x_143, x_145, x_146, x_145, x_147, x_148, x_2, x_3, x_4, x_5, x_135); -lean_dec(x_147); -lean_dec(x_145); -lean_dec(x_143); -if (lean_obj_tag(x_149) == 0) -{ -uint8_t x_150; -x_150 = !lean_is_exclusive(x_149); -if (x_150 == 0) -{ -lean_object* x_151; -x_151 = lean_ctor_get(x_149, 0); -lean_dec(x_151); -lean_ctor_set(x_149, 0, x_148); -return x_149; -} -else -{ -lean_object* x_152; lean_object* x_153; -x_152 = lean_ctor_get(x_149, 1); -lean_inc(x_152); -lean_dec(x_149); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_148); -lean_ctor_set(x_153, 1, x_152); -return x_153; -} +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_45, 1); +lean_inc(x_51); +lean_dec(x_45); +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 -{ -uint8_t x_154; -x_154 = !lean_is_exclusive(x_149); -if (x_154 == 0) -{ -return x_149; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_149, 0); -x_156 = lean_ctor_get(x_149, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_149); -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; -} -} +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_45, 1); +lean_inc(x_54); +lean_dec(x_45); +x_55 = lean_box(0); +x_56 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__2(x_1, x_44, x_55, x_2, x_3, x_4, x_5, x_6, x_54); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_56; } } else { -uint8_t x_158; +uint8_t x_57; +lean_dec(x_44); +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_158 = !lean_is_exclusive(x_131); -if (x_158 == 0) +x_57 = !lean_is_exclusive(x_45); +if (x_57 == 0) { -return x_131; +return x_45; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_131, 0); -x_160 = lean_ctor_get(x_131, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_131); -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_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_45, 0); +x_59 = lean_ctor_get(x_45, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_45); +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; } -case 5: -{ -lean_object* x_162; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_162 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_162) == 0) -{ -lean_object* x_163; -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -if (lean_obj_tag(x_163) == 0) -{ -lean_object* x_164; lean_object* x_165; -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -lean_dec(x_162); -x_165 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_164); -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_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_166 = lean_ctor_get(x_162, 1); -lean_inc(x_166); -lean_dec(x_162); -x_167 = lean_ctor_get(x_163, 0); -lean_inc(x_167); -lean_dec(x_163); -x_168 = lean_unsigned_to_nat(0u); -x_169 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_168); -x_170 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_169); -x_171 = lean_mk_array(x_169, x_170); -x_172 = lean_unsigned_to_nat(1u); -x_173 = lean_nat_sub(x_169, x_172); -lean_dec(x_169); -x_174 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_171, x_173); -x_175 = lean_ctor_get(x_167, 3); -lean_inc(x_175); -lean_dec(x_167); -x_176 = lean_ctor_get(x_175, 1); -lean_inc(x_176); -x_177 = lean_ctor_get(x_175, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_175, 2); -lean_inc(x_178); -lean_dec(x_175); -x_179 = lean_box(0); -lean_inc(x_176); -x_180 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__6(x_174, x_176, x_177, x_176, x_178, x_179, x_2, x_3, x_4, x_5, x_166); -lean_dec(x_178); -lean_dec(x_176); -lean_dec(x_174); -if (lean_obj_tag(x_180) == 0) -{ -uint8_t x_181; -x_181 = !lean_is_exclusive(x_180); -if (x_181 == 0) -{ -lean_object* x_182; -x_182 = lean_ctor_get(x_180, 0); -lean_dec(x_182); -lean_ctor_set(x_180, 0, x_179); -return x_180; } -else -{ -lean_object* x_183; lean_object* x_184; -x_183 = lean_ctor_get(x_180, 1); -lean_inc(x_183); -lean_dec(x_180); -x_184 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_184, 0, x_179); -lean_ctor_set(x_184, 1, x_183); -return x_184; } } -else -{ -uint8_t x_185; -x_185 = !lean_is_exclusive(x_180); -if (x_185 == 0) -{ -return x_180; } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_186 = lean_ctor_get(x_180, 0); -x_187 = lean_ctor_get(x_180, 1); -lean_inc(x_187); -lean_inc(x_186); -lean_dec(x_180); -x_188 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_188, 0, x_186); -lean_ctor_set(x_188, 1, x_187); -return x_188; +lean_object* x_10; +x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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_Compiler_Simp_inlineCandidate_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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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: { -uint8_t x_189; +lean_object* x_11; +x_11 = l_Lean_Compiler_Simp_inlineCandidate_x3f___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_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_189 = !lean_is_exclusive(x_162); -if (x_189 == 0) -{ -return x_162; +return x_11; } -else -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_190 = lean_ctor_get(x_162, 0); -x_191 = lean_ctor_get(x_162, 1); -lean_inc(x_191); -lean_inc(x_190); -lean_dec(x_162); -x_192 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_192, 0, x_190); -lean_ctor_set(x_192, 1, x_191); -return x_192; } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_Compiler_Simp_inlineCandidate_x3f___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_7); +lean_dec(x_6); +lean_dec(x_5); +return x_11; } } -case 6: +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___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_193; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_193 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_193) == 0) +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_Lean_Compiler_Simp_expandTrivialExpr(lean_object* x_1, 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_194; -x_194 = lean_ctor_get(x_193, 0); -lean_inc(x_194); -if (lean_obj_tag(x_194) == 0) +uint8_t x_8; +x_8 = l_Lean_Expr_isFVar(x_1); +if (x_8 == 0) { -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_193, 1); -lean_inc(x_195); -lean_dec(x_193); -x_196 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_195); -return x_196; +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_7); +return x_9; } 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_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_197 = lean_ctor_get(x_193, 1); -lean_inc(x_197); -lean_dec(x_193); -x_198 = lean_ctor_get(x_194, 0); -lean_inc(x_198); -lean_dec(x_194); -x_199 = lean_unsigned_to_nat(0u); -x_200 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_199); -x_201 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_200); -x_202 = lean_mk_array(x_200, x_201); -x_203 = lean_unsigned_to_nat(1u); -x_204 = lean_nat_sub(x_200, x_203); -lean_dec(x_200); -x_205 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_202, x_204); -x_206 = lean_ctor_get(x_198, 3); -lean_inc(x_206); -lean_dec(x_198); -x_207 = lean_ctor_get(x_206, 1); -lean_inc(x_207); -x_208 = lean_ctor_get(x_206, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_206, 2); -lean_inc(x_209); -lean_dec(x_206); -x_210 = lean_box(0); -lean_inc(x_207); -x_211 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__7(x_205, x_207, x_208, x_207, x_209, x_210, x_2, x_3, x_4, x_5, x_197); -lean_dec(x_209); -lean_dec(x_207); -lean_dec(x_205); -if (lean_obj_tag(x_211) == 0) -{ -uint8_t x_212; -x_212 = !lean_is_exclusive(x_211); -if (x_212 == 0) +uint8_t x_10; lean_object* x_11; uint8_t x_12; +x_10 = 1; +lean_inc(x_1); +x_11 = l_Lean_Compiler_Simp_findExpr(x_1, x_10, x_4, x_5, x_6, x_7); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_213; -x_213 = lean_ctor_get(x_211, 0); -lean_dec(x_213); -lean_ctor_set(x_211, 0, x_210); -return x_211; -} -else +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = l_Lean_Expr_isLambda(x_13); +if (x_15 == 0) { -lean_object* x_214; lean_object* x_215; -x_214 = lean_ctor_get(x_211, 1); -lean_inc(x_214); -lean_dec(x_211); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_210); -lean_ctor_set(x_215, 1, x_214); -return x_215; -} -} -else +uint8_t x_16; +x_16 = lean_expr_eqv(x_1, x_13); +if (x_16 == 0) { -uint8_t x_216; -x_216 = !lean_is_exclusive(x_211); -if (x_216 == 0) +lean_object* x_17; uint8_t x_18; +lean_free_object(x_11); +lean_dec(x_1); +x_17 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -return x_211; +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +lean_ctor_set(x_17, 0, x_13); +return x_17; } else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_217 = lean_ctor_get(x_211, 0); -x_218 = lean_ctor_get(x_211, 1); -lean_inc(x_218); -lean_inc(x_217); -lean_dec(x_211); -x_219 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -return x_219; -} -} +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_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } else { -uint8_t x_220; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_220 = !lean_is_exclusive(x_193); -if (x_220 == 0) -{ -return x_193; +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_1); +return x_11; +} } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_221 = lean_ctor_get(x_193, 0); -x_222 = lean_ctor_get(x_193, 1); -lean_inc(x_222); -lean_inc(x_221); -lean_dec(x_193); -x_223 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_223, 0, x_221); -lean_ctor_set(x_223, 1, x_222); -return x_223; -} -} +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_1); +return x_11; } -case 7: -{ -lean_object* x_224; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_224 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_224) == 0) -{ -lean_object* x_225; -x_225 = lean_ctor_get(x_224, 0); -lean_inc(x_225); -if (lean_obj_tag(x_225) == 0) -{ -lean_object* x_226; lean_object* x_227; -x_226 = lean_ctor_get(x_224, 1); -lean_inc(x_226); -lean_dec(x_224); -x_227 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_226); -return x_227; } else { -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; -x_228 = lean_ctor_get(x_224, 1); -lean_inc(x_228); -lean_dec(x_224); -x_229 = lean_ctor_get(x_225, 0); -lean_inc(x_229); -lean_dec(x_225); -x_230 = lean_unsigned_to_nat(0u); -x_231 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_230); -x_232 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_231); -x_233 = lean_mk_array(x_231, x_232); -x_234 = lean_unsigned_to_nat(1u); -x_235 = lean_nat_sub(x_231, x_234); -lean_dec(x_231); -x_236 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_233, x_235); -x_237 = lean_ctor_get(x_229, 3); -lean_inc(x_237); -lean_dec(x_229); -x_238 = lean_ctor_get(x_237, 1); -lean_inc(x_238); -x_239 = lean_ctor_get(x_237, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_237, 2); -lean_inc(x_240); -lean_dec(x_237); -x_241 = lean_box(0); -lean_inc(x_238); -x_242 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__8(x_236, x_238, x_239, x_238, x_240, x_241, x_2, x_3, x_4, x_5, x_228); -lean_dec(x_240); -lean_dec(x_238); -lean_dec(x_236); -if (lean_obj_tag(x_242) == 0) +lean_object* x_22; lean_object* x_23; uint8_t 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 = l_Lean_Expr_isLambda(x_22); +if (x_24 == 0) { -uint8_t x_243; -x_243 = !lean_is_exclusive(x_242); -if (x_243 == 0) +uint8_t x_25; +x_25 = lean_expr_eqv(x_1, x_22); +if (x_25 == 0) { -lean_object* x_244; -x_244 = lean_ctor_get(x_242, 0); -lean_dec(x_244); -lean_ctor_set(x_242, 0, x_241); -return x_242; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_26 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_23); +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); } -else -{ -lean_object* x_245; lean_object* x_246; -x_245 = lean_ctor_get(x_242, 1); -lean_inc(x_245); -lean_dec(x_242); -x_246 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_246, 0, x_241); -lean_ctor_set(x_246, 1, x_245); -return x_246; +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; } +lean_ctor_set(x_29, 0, x_22); +lean_ctor_set(x_29, 1, x_27); +return x_29; } else { -uint8_t x_247; -x_247 = !lean_is_exclusive(x_242); -if (x_247 == 0) -{ -return x_242; +lean_object* x_30; +lean_dec(x_22); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_23); +return x_30; +} } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_248 = lean_ctor_get(x_242, 0); -x_249 = lean_ctor_get(x_242, 1); -lean_inc(x_249); -lean_inc(x_248); -lean_dec(x_242); -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; +lean_object* x_31; +lean_dec(x_22); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_23); +return x_31; } } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___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: { -uint8_t x_251; +lean_object* x_9; +x_9 = l_Lean_Compiler_Simp_expandTrivialExpr___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); -x_251 = !lean_is_exclusive(x_224); -if (x_251 == 0) -{ -return x_224; -} -else -{ -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_224, 0); -x_253 = lean_ctor_get(x_224, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_224); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -return x_254; -} -} -} -case 8: -{ -lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_255 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_collectInlineStats_go___lambda__2), 6, 1); -lean_closure_set(x_255, 0, x_1); -x_256 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__2; -x_257 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__1___rarg), 7, 2); -lean_closure_set(x_257, 0, x_255); -lean_closure_set(x_257, 1, x_256); -x_258 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_collectInlineStats_goLambda___spec__2(x_257, x_2, x_3, x_4, x_5, x_6); -return x_258; -} -case 9: -{ -lean_object* x_259; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_259 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_259) == 0) -{ -lean_object* x_260; -x_260 = lean_ctor_get(x_259, 0); -lean_inc(x_260); -if (lean_obj_tag(x_260) == 0) -{ -lean_object* x_261; lean_object* x_262; -x_261 = lean_ctor_get(x_259, 1); -lean_inc(x_261); -lean_dec(x_259); -x_262 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_261); -return x_262; -} -else -{ -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; -x_263 = lean_ctor_get(x_259, 1); -lean_inc(x_263); -lean_dec(x_259); -x_264 = lean_ctor_get(x_260, 0); -lean_inc(x_264); -lean_dec(x_260); -x_265 = lean_unsigned_to_nat(0u); -x_266 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_265); -x_267 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_266); -x_268 = lean_mk_array(x_266, x_267); -x_269 = lean_unsigned_to_nat(1u); -x_270 = lean_nat_sub(x_266, x_269); -lean_dec(x_266); -x_271 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_268, x_270); -x_272 = lean_ctor_get(x_264, 3); -lean_inc(x_272); -lean_dec(x_264); -x_273 = lean_ctor_get(x_272, 1); -lean_inc(x_273); -x_274 = lean_ctor_get(x_272, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_272, 2); -lean_inc(x_275); -lean_dec(x_272); -x_276 = lean_box(0); -lean_inc(x_273); -x_277 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__9(x_271, x_273, x_274, x_273, x_275, x_276, x_2, x_3, x_4, x_5, x_263); -lean_dec(x_275); -lean_dec(x_273); -lean_dec(x_271); -if (lean_obj_tag(x_277) == 0) -{ -uint8_t x_278; -x_278 = !lean_is_exclusive(x_277); -if (x_278 == 0) -{ -lean_object* x_279; -x_279 = lean_ctor_get(x_277, 0); -lean_dec(x_279); -lean_ctor_set(x_277, 0, x_276); -return x_277; -} -else -{ -lean_object* x_280; lean_object* x_281; -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -lean_dec(x_277); -x_281 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_281, 0, x_276); -lean_ctor_set(x_281, 1, x_280); -return x_281; -} -} -else -{ -uint8_t x_282; -x_282 = !lean_is_exclusive(x_277); -if (x_282 == 0) -{ -return x_277; -} -else -{ -lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_283 = lean_ctor_get(x_277, 0); -x_284 = lean_ctor_get(x_277, 1); -lean_inc(x_284); -lean_inc(x_283); -lean_dec(x_277); -x_285 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_285, 0, x_283); -lean_ctor_set(x_285, 1, x_284); -return x_285; -} -} +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___boxed(lean_object* x_1, 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_286; +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_expandTrivialExpr(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); -x_286 = !lean_is_exclusive(x_259); -if (x_286 == 0) -{ -return x_259; -} -else -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_287 = lean_ctor_get(x_259, 0); -x_288 = lean_ctor_get(x_259, 1); -lean_inc(x_288); -lean_inc(x_287); -lean_dec(x_259); -x_289 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_289, 0, x_287); -lean_ctor_set(x_289, 1, x_288); -return x_289; -} +return x_8; } } -case 10: -{ -lean_object* x_290; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_290 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_290) == 0) +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_291; -x_291 = lean_ctor_get(x_290, 0); -lean_inc(x_291); -if (lean_obj_tag(x_291) == 0) +if (lean_obj_tag(x_5) == 8) { -lean_object* x_292; lean_object* x_293; -x_292 = lean_ctor_get(x_290, 1); -lean_inc(x_292); -lean_dec(x_290); -x_293 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_292); -return x_293; -} -else -{ -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; -x_294 = lean_ctor_get(x_290, 1); -lean_inc(x_294); -lean_dec(x_290); -x_295 = lean_ctor_get(x_291, 0); -lean_inc(x_295); -lean_dec(x_291); -x_296 = lean_unsigned_to_nat(0u); -x_297 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_296); -x_298 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_297); -x_299 = lean_mk_array(x_297, x_298); -x_300 = lean_unsigned_to_nat(1u); -x_301 = lean_nat_sub(x_297, x_300); -lean_dec(x_297); -x_302 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_299, x_301); -x_303 = lean_ctor_get(x_295, 3); -lean_inc(x_303); -lean_dec(x_295); -x_304 = lean_ctor_get(x_303, 1); -lean_inc(x_304); -x_305 = lean_ctor_get(x_303, 0); -lean_inc(x_305); -x_306 = lean_ctor_get(x_303, 2); -lean_inc(x_306); -lean_dec(x_303); -x_307 = lean_box(0); -lean_inc(x_304); -x_308 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__10(x_302, x_304, x_305, x_304, x_306, x_307, x_2, x_3, x_4, x_5, x_294); -lean_dec(x_306); -lean_dec(x_304); -lean_dec(x_302); -if (lean_obj_tag(x_308) == 0) -{ -uint8_t x_309; -x_309 = !lean_is_exclusive(x_308); -if (x_309 == 0) -{ -lean_object* x_310; -x_310 = lean_ctor_get(x_308, 0); -lean_dec(x_310); -lean_ctor_set(x_308, 0, x_307); -return x_308; +lean_object* x_7; lean_object* 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; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_5, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_5, 3); +lean_inc(x_10); +x_11 = lean_ctor_get_uint8(x_5, sizeof(void*)*4 + 8); +lean_dec(x_5); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_6, x_12); +x_14 = l_Lean_Compiler_Simp_mkFlatLet_go(x_1, x_2, x_3, x_4, x_10, x_13); +lean_dec(x_13); +x_15 = l_Lean_Expr_letE___override(x_7, x_8, x_9, x_14, x_11); +return x_15; } else { -lean_object* x_311; lean_object* x_312; -x_311 = lean_ctor_get(x_308, 1); -lean_inc(x_311); -lean_dec(x_308); -x_312 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_312, 0, x_307); -lean_ctor_set(x_312, 1, x_311); -return x_312; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_expr_lift_loose_bvars(x_3, x_16, x_6); +x_18 = l_Lean_Expr_letE___override(x_1, x_2, x_5, x_17, x_4); +return x_18; } } -else -{ -uint8_t x_313; -x_313 = !lean_is_exclusive(x_308); -if (x_313 == 0) -{ -return x_308; } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet_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_314; lean_object* x_315; lean_object* x_316; -x_314 = lean_ctor_get(x_308, 0); -x_315 = lean_ctor_get(x_308, 1); -lean_inc(x_315); -lean_inc(x_314); -lean_dec(x_308); -x_316 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_316, 0, x_314); -lean_ctor_set(x_316, 1, x_315); -return x_316; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_4); +lean_dec(x_4); +x_8 = l_Lean_Compiler_Simp_mkFlatLet_go(x_1, x_2, x_3, x_7, x_5, x_6); +lean_dec(x_6); +lean_dec(x_3); +return x_8; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Lean_Compiler_Simp_mkFlatLet_go(x_1, x_2, x_4, x_5, x_3, x_6); +return x_7; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_mkFlatLet___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_317; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_5); lean_dec(x_5); +x_7 = l_Lean_Compiler_Simp_mkFlatLet(x_1, x_2, x_3, x_4, x_6); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_317 = !lean_is_exclusive(x_290); -if (x_317 == 0) -{ -return x_290; -} -else -{ -lean_object* x_318; lean_object* x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_290, 0); -x_319 = lean_ctor_get(x_290, 1); -lean_inc(x_319); -lean_inc(x_318); -lean_dec(x_290); -x_320 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_320, 0, x_318); -lean_ctor_set(x_320, 1, x_319); -return x_320; -} +return x_7; } } -default: -{ -lean_object* x_321; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_321 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_4, x_5, x_6); -if (lean_obj_tag(x_321) == 0) -{ -lean_object* x_322; -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -if (lean_obj_tag(x_322) == 0) +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1() { +_start: { -lean_object* x_323; lean_object* x_324; -x_323 = lean_ctor_get(x_321, 1); -lean_inc(x_323); -lean_dec(x_321); -x_324 = l_Lean_Compiler_Simp_collectInlineStats_goValue(x_1, x_2, x_3, x_4, x_5, x_323); -return x_324; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_instMonadCoreM; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } -else -{ -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; -x_325 = lean_ctor_get(x_321, 1); -lean_inc(x_325); -lean_dec(x_321); -x_326 = lean_ctor_get(x_322, 0); -lean_inc(x_326); -lean_dec(x_322); -x_327 = lean_unsigned_to_nat(0u); -x_328 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_327); -x_329 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_328); -x_330 = lean_mk_array(x_328, x_329); -x_331 = lean_unsigned_to_nat(1u); -x_332 = lean_nat_sub(x_328, x_331); -lean_dec(x_328); -x_333 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_330, x_332); -x_334 = lean_ctor_get(x_326, 3); -lean_inc(x_334); -lean_dec(x_326); -x_335 = lean_ctor_get(x_334, 1); -lean_inc(x_335); -x_336 = lean_ctor_get(x_334, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_334, 2); -lean_inc(x_337); -lean_dec(x_334); -x_338 = lean_box(0); -lean_inc(x_335); -x_339 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__11(x_333, x_335, x_336, x_335, x_337, x_338, x_2, x_3, x_4, x_5, x_325); -lean_dec(x_337); -lean_dec(x_335); -lean_dec(x_333); -if (lean_obj_tag(x_339) == 0) -{ -uint8_t x_340; -x_340 = !lean_is_exclusive(x_339); -if (x_340 == 0) -{ -lean_object* x_341; -x_341 = lean_ctor_get(x_339, 0); -lean_dec(x_341); -lean_ctor_set(x_339, 0, x_338); -return x_339; } -else +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2() { +_start: { -lean_object* x_342; lean_object* x_343; -x_342 = lean_ctor_get(x_339, 1); -lean_inc(x_342); -lean_dec(x_339); -x_343 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_343, 0, x_338); -lean_ctor_set(x_343, 1, x_342); -return x_343; -} +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } -else -{ -uint8_t x_344; -x_344 = !lean_is_exclusive(x_339); -if (x_344 == 0) -{ -return x_339; } -else +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3() { +_start: { -lean_object* x_345; lean_object* x_346; lean_object* x_347; -x_345 = lean_ctor_get(x_339, 0); -x_346 = lean_ctor_get(x_339, 1); -lean_inc(x_346); -lean_inc(x_345); -lean_dec(x_339); -x_347 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_347, 0, x_345); -lean_ctor_set(x_347, 1, x_346); -return x_347; -} -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } -else -{ -uint8_t x_348; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_348 = !lean_is_exclusive(x_321); -if (x_348 == 0) -{ -return x_321; } -else +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4() { +_start: { -lean_object* x_349; lean_object* x_350; lean_object* x_351; -x_349 = lean_ctor_get(x_321, 0); -x_350 = lean_ctor_get(x_321, 1); -lean_inc(x_350); -lean_inc(x_349); -lean_dec(x_321); -x_351 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_351, 0, x_349); -lean_ctor_set(x_351, 1, x_350); -return x_351; -} -} -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3; +x_2 = l_OptionT_instMonadOptionT___rarg(x_1); +return x_2; } } -static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1() { +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5() { _start: { -lean_object* x_1; -x_1 = l_Lean_inheritedTraceOptions; -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___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_EXPORT lean_object* l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___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_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_st_ref_get(x_5, x_6); -x_8 = lean_ctor_get(x_7, 1); -lean_inc(x_8); -lean_dec(x_7); -x_9 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1; -x_10 = lean_st_ref_get(x_9, x_8); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_4, 2); -x_14 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_13, x_1); -lean_dec(x_12); -x_15 = lean_box(x_14); -lean_ctor_set(x_10, 0, x_15); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; } -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -x_18 = lean_ctor_get(x_4, 2); -x_19 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_16, x_18, x_1); -lean_dec(x_16); -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_17); -return x_21; -} } -} -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1() { _start: { lean_object* x_1; -x_1 = l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp", 18); return x_1; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp.inlineProjInst?.visitProj", 44); +return x_1; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_object* x_1; +x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); +return x_1; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_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_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1; +x_2 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2; +x_3 = lean_unsigned_to_nat(272u); +x_4 = lean_unsigned_to_nat(27u); +x_5 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj(lean_object* x_1, 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_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6() { -_start: +if (lean_obj_tag(x_1) == 11) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3; -x_3 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4; -x_4 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5; -x_5 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_1); -lean_ctor_set(x_5, 2, x_2); -lean_ctor_set(x_5, 3, x_3); -lean_ctor_set(x_5, 4, x_4); -lean_ctor_set(x_5, 5, x_2); -lean_ctor_set(x_5, 6, x_3); -lean_ctor_set(x_5, 7, x_3); -return x_5; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +lean_dec(x_1); +lean_inc(x_6); +x_10 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visit(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +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_6); +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; } -static lean_object* _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7() { -_start: +else { -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_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; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___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: +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_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; uint8_t x_26; -x_8 = lean_ctor_get(x_5, 5); -x_9 = lean_st_ref_get(x_6, 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_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_st_ref_get(x_6, x_11); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_st_ref_get(x_4, x_14); -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_ctor_get(x_16, 0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_10, 1); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_5, 2); -x_20 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6; +lean_dec(x_10); +x_19 = lean_ctor_get(x_11, 0); lean_inc(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_20); -lean_ctor_set(x_21, 2, x_18); -lean_ctor_set(x_21, 3, x_19); -x_22 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_2); -x_23 = lean_st_ref_take(x_6, x_17); -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_dec(x_11); +x_20 = lean_st_ref_get(x_6, x_18); +lean_dec(x_6); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Expr_constructorApp_x3f(x_23, x_19); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; +lean_dec(x_8); +x_25 = lean_box(0); +lean_ctor_set(x_20, 0, x_25); +return x_20; +} +else +{ +uint8_t x_26; x_26 = !lean_is_exclusive(x_24); 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; uint8_t x_34; -x_27 = lean_ctor_get(x_24, 3); -x_28 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_29 = 0; -x_30 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_30, 0, x_1); -lean_ctor_set(x_30, 1, x_22); -lean_ctor_set(x_30, 2, x_28); -lean_ctor_set_uint8(x_30, sizeof(void*)*3, x_29); -lean_inc(x_8); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_8); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Std_PersistentArray_push___rarg(x_27, x_31); -lean_ctor_set(x_24, 3, x_32); -x_33 = lean_st_ref_set(x_6, x_24, x_25); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 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; uint8_t x_33; +x_27 = lean_ctor_get(x_24, 0); +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, 3); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_nat_add(x_30, x_8); +lean_dec(x_8); +lean_dec(x_30); +x_32 = lean_array_get_size(x_29); +x_33 = lean_nat_dec_lt(x_31, x_32); +lean_dec(x_32); +if (x_33 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_object* x_34; lean_object* x_35; +lean_dec(x_31); +lean_dec(x_29); +x_34 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_35 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_34); +lean_ctor_set(x_24, 0, x_35); +lean_ctor_set(x_20, 0, x_24); +return x_20; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -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; +lean_object* x_36; +x_36 = lean_array_fget(x_29, x_31); +lean_dec(x_31); +lean_dec(x_29); +lean_ctor_set(x_24, 0, x_36); +lean_ctor_set(x_20, 0, x_24); +return x_20; } } 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; 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; lean_object* x_56; lean_object* x_57; -x_40 = lean_ctor_get(x_24, 0); -x_41 = lean_ctor_get(x_24, 1); -x_42 = lean_ctor_get(x_24, 2); -x_43 = lean_ctor_get(x_24, 3); -x_44 = lean_ctor_get(x_24, 4); -x_45 = lean_ctor_get(x_24, 5); -x_46 = lean_ctor_get(x_24, 6); -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_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; +x_37 = lean_ctor_get(x_24, 0); +lean_inc(x_37); lean_dec(x_24); -x_47 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_48 = 0; -x_49 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_22); -lean_ctor_set(x_49, 2, x_47); -lean_ctor_set_uint8(x_49, sizeof(void*)*3, x_48); -lean_inc(x_8); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_8); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Std_PersistentArray_push___rarg(x_43, x_50); -x_52 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_52, 0, x_40); -lean_ctor_set(x_52, 1, x_41); -lean_ctor_set(x_52, 2, x_42); -lean_ctor_set(x_52, 3, x_51); -lean_ctor_set(x_52, 4, x_44); -lean_ctor_set(x_52, 5, x_45); -lean_ctor_set(x_52, 6, x_46); -x_53 = lean_st_ref_set(x_6, x_52, x_25); -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; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 3); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_nat_add(x_40, x_8); +lean_dec(x_8); +lean_dec(x_40); +x_42 = lean_array_get_size(x_39); +x_43 = lean_nat_dec_lt(x_41, x_42); +lean_dec(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_41); +lean_dec(x_39); +x_44 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_45 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_44); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_20, 0, x_46); +return x_20; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_array_fget(x_39, x_41); +lean_dec(x_41); +lean_dec(x_39); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_20, 0, x_48); +return x_20; +} +} +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_ctor_get(x_20, 0); +x_50 = lean_ctor_get(x_20, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_20); +x_51 = lean_ctor_get(x_49, 0); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Expr_constructorApp_x3f(x_51, x_19); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_8); +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_50); +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; uint8_t x_62; +x_55 = lean_ctor_get(x_52, 0); +lean_inc(x_55); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + x_56 = x_52; } else { - lean_dec_ref(x_53); - x_55 = lean_box(0); + lean_dec_ref(x_52); + x_56 = lean_box(0); } -x_56 = lean_box(0); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); +x_57 = lean_ctor_get(x_55, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_ctor_get(x_57, 3); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_nat_add(x_59, x_8); +lean_dec(x_8); +lean_dec(x_59); +x_61 = lean_array_get_size(x_58); +x_62 = lean_nat_dec_lt(x_60, x_61); +lean_dec(x_61); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_60); +lean_dec(x_58); +x_63 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_64 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_63); +if (lean_is_scalar(x_56)) { + x_65 = lean_alloc_ctor(1, 1, 0); } else { - x_57 = x_55; + x_65 = x_56; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; +lean_ctor_set(x_65, 0, 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_50); +return x_66; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_array_fget(x_58, x_60); +lean_dec(x_60); +lean_dec(x_58); +if (lean_is_scalar(x_56)) { + x_68 = lean_alloc_ctor(1, 1, 0); +} else { + x_68 = x_56; } +lean_ctor_set(x_68, 0, x_67); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_50); +return x_69; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___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 = l_Lean_LocalDecl_value(x_1); -x_9 = l_Lean_Expr_isLambda(x_8); -if (x_9 == 0) +} +} +} +else { -lean_object* x_10; lean_object* x_11; +uint8_t x_70; lean_dec(x_8); lean_dec(x_6); -lean_dec(x_5); -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_7); -return x_11; +x_70 = !lean_is_exclusive(x_10); +if (x_70 == 0) +{ +return x_10; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_10, 0); +x_72 = lean_ctor_get(x_10, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_10); +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; lean_object* x_75; +lean_dec(x_1); +x_74 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4; +x_75 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1(x_74, x_2, x_3, x_4, x_5, x_6, x_7); +return x_75; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___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_1) == 8) +{ +lean_object* x_7; lean_object* 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; lean_object* x_17; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_dec(x_1); +x_12 = lean_expr_instantiate_rev(x_8, x_2); +lean_dec(x_8); +x_13 = lean_expr_instantiate_rev(x_9, x_2); +lean_dec(x_9); +x_14 = l_Lean_Compiler_mkLetDecl(x_7, x_12, x_13, x_11, x_3, x_4, x_5, x_6); +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_array_push(x_2, x_15); +x_1 = x_10; +x_2 = x_17; +x_6 = x_16; +goto _start; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_expr_instantiate_rev(x_1, x_2); +lean_dec(x_2); +lean_dec(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_6); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineProjInst_x3f_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: +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = 1; +x_9 = l_Lean_Compiler_Simp_findExpr(x_1, x_8, x_4, x_5, x_6, 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_st_ref_get(x_6, x_11); +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_12, 1); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Expr_isConstructorApp(x_16, x_10); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = l_Lean_Expr_isProj(x_10); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = l_Lean_Expr_getAppFn(x_10); +if (lean_obj_tag(x_19) == 4) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_12); +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_inc(x_6); +lean_inc(x_5); +x_22 = l_Lean_Compiler_getStage1Decl_x3f(x_20, x_5, x_6, x_15); +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_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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 +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_22); +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; uint8_t x_37; +x_31 = lean_ctor_get(x_22, 1); +x_32 = lean_ctor_get(x_22, 0); +lean_dec(x_32); +x_33 = lean_ctor_get(x_23, 0); +lean_inc(x_33); +lean_dec(x_23); +x_34 = l_Lean_Compiler_Decl_getArity(x_33); +x_35 = lean_unsigned_to_nat(0u); +x_36 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_10, x_35); +x_37 = lean_nat_dec_eq(x_34, x_36); +lean_dec(x_34); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_38 = lean_box(0); +lean_ctor_set(x_22, 0, x_38); +return x_22; +} +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_free_object(x_22); +x_39 = lean_ctor_get(x_33, 3); +lean_inc(x_39); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_dec(x_33); +x_41 = l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(x_40, x_21, x_39); +lean_dec(x_21); +lean_dec(x_40); +x_42 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_36); +x_43 = lean_mk_array(x_36, x_42); +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_sub(x_36, x_44); +lean_dec(x_36); +x_46 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_10, x_43, x_45); +x_47 = l_Lean_Expr_beta(x_41, x_46); +x_48 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_49 = l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__1(x_47, x_48, x_4, x_5, x_6, x_31); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_1 = x_50; +x_7 = x_51; +goto _start; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_53 = lean_ctor_get(x_22, 1); +lean_inc(x_53); +lean_dec(x_22); +x_54 = lean_ctor_get(x_23, 0); +lean_inc(x_54); +lean_dec(x_23); +x_55 = l_Lean_Compiler_Decl_getArity(x_54); +x_56 = lean_unsigned_to_nat(0u); +x_57 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_10, x_56); +x_58 = lean_nat_dec_eq(x_55, x_57); +lean_dec(x_55); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_57); +lean_dec(x_54); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_53); +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; +x_61 = lean_ctor_get(x_54, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_63 = l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(x_62, x_21, x_61); +lean_dec(x_21); +lean_dec(x_62); +x_64 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_57); +x_65 = lean_mk_array(x_57, x_64); +x_66 = lean_unsigned_to_nat(1u); +x_67 = lean_nat_sub(x_57, x_66); +lean_dec(x_57); +x_68 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_10, x_65, x_67); +x_69 = l_Lean_Expr_beta(x_63, x_68); +x_70 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_71 = l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__1(x_69, x_70, x_4, x_5, x_6, x_53); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_1 = x_72; +x_7 = x_73; +goto _start; +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_75 = !lean_is_exclusive(x_22); +if (x_75 == 0) +{ +return x_22; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_22, 0); +x_77 = lean_ctor_get(x_22, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_22); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +else +{ +lean_object* x_79; +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_79 = lean_box(0); +lean_ctor_set(x_12, 0, x_79); +return x_12; +} +} +else +{ +lean_object* x_80; +lean_free_object(x_12); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_80 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj(x_10, x_2, x_3, x_4, x_5, x_6, x_15); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) +{ +uint8_t x_82; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_82 = !lean_is_exclusive(x_80); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_80, 0); +lean_dec(x_83); +x_84 = lean_box(0); +lean_ctor_set(x_80, 0, x_84); +return x_80; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_80, 1); +lean_inc(x_85); +lean_dec(x_80); +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); +return x_87; +} +} +else +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_80, 1); +lean_inc(x_88); +lean_dec(x_80); +x_89 = lean_ctor_get(x_81, 0); +lean_inc(x_89); +lean_dec(x_81); +x_1 = x_89; +x_7 = x_88; +goto _start; +} +} +else +{ +uint8_t x_91; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_91 = !lean_is_exclusive(x_80); +if (x_91 == 0) +{ +return x_80; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_80, 0); +x_93 = lean_ctor_get(x_80, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_80); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +} +else +{ +lean_object* x_95; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_10); +lean_ctor_set(x_12, 0, x_95); +return x_12; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_96 = lean_ctor_get(x_12, 0); +x_97 = lean_ctor_get(x_12, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_12); +x_98 = lean_ctor_get(x_96, 0); +lean_inc(x_98); +lean_dec(x_96); +x_99 = l_Lean_Expr_isConstructorApp(x_98, x_10); +if (x_99 == 0) +{ +uint8_t x_100; +x_100 = l_Lean_Expr_isProj(x_10); +if (x_100 == 0) +{ +lean_object* x_101; +x_101 = l_Lean_Expr_getAppFn(x_10); +if (lean_obj_tag(x_101) == 4) +{ +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); +lean_inc(x_6); +lean_inc(x_5); +x_104 = l_Lean_Compiler_getStage1Decl_x3f(x_102, x_5, x_6, x_97); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_103); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; +} else { + lean_dec_ref(x_104); + x_107 = lean_box(0); +} +x_108 = lean_box(0); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); +} else { + x_109 = x_107; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; +} +else +{ +lean_object* x_110; 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_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_111 = x_104; +} else { + lean_dec_ref(x_104); + x_111 = lean_box(0); +} +x_112 = lean_ctor_get(x_105, 0); +lean_inc(x_112); +lean_dec(x_105); +x_113 = l_Lean_Compiler_Decl_getArity(x_112); +x_114 = lean_unsigned_to_nat(0u); +x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_10, x_114); +x_116 = lean_nat_dec_eq(x_113, x_115); +lean_dec(x_113); +if (x_116 == 0) +{ +lean_object* x_117; lean_object* x_118; +lean_dec(x_115); +lean_dec(x_112); +lean_dec(x_103); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_117 = lean_box(0); +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_117); +lean_ctor_set(x_118, 1, x_110); +return x_118; +} +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; 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_119 = lean_ctor_get(x_112, 3); +lean_inc(x_119); +x_120 = lean_ctor_get(x_112, 1); +lean_inc(x_120); +lean_dec(x_112); +x_121 = l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(x_120, x_103, x_119); +lean_dec(x_103); +lean_dec(x_120); +x_122 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_115); +x_123 = lean_mk_array(x_115, x_122); +x_124 = lean_unsigned_to_nat(1u); +x_125 = lean_nat_sub(x_115, x_124); +lean_dec(x_115); +x_126 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_10, x_123, x_125); +x_127 = l_Lean_Expr_beta(x_121, x_126); +x_128 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_129 = l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__1(x_127, x_128, x_4, x_5, x_6, x_110); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +lean_dec(x_129); +x_1 = x_130; +x_7 = x_131; +goto _start; +} +} +} +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_dec(x_103); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_133 = lean_ctor_get(x_104, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_104, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_135 = x_104; +} else { + lean_dec_ref(x_104); + x_135 = lean_box(0); +} +if (lean_is_scalar(x_135)) { + x_136 = lean_alloc_ctor(1, 2, 0); +} else { + x_136 = x_135; +} +lean_ctor_set(x_136, 0, x_133); +lean_ctor_set(x_136, 1, x_134); +return x_136; +} +} +else +{ +lean_object* x_137; lean_object* x_138; +lean_dec(x_101); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_137 = lean_box(0); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_97); +return x_138; +} +} +else +{ +lean_object* x_139; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_139 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj(x_10, x_2, x_3, x_4, x_5, x_6, x_97); +if (lean_obj_tag(x_139) == 0) +{ +lean_object* x_140; +x_140 = lean_ctor_get(x_139, 0); +lean_inc(x_140); +if (lean_obj_tag(x_140) == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_141 = lean_ctor_get(x_139, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_142 = x_139; +} else { + lean_dec_ref(x_139); + 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(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_141); +return x_144; +} +else +{ +lean_object* x_145; lean_object* x_146; +x_145 = lean_ctor_get(x_139, 1); +lean_inc(x_145); +lean_dec(x_139); +x_146 = lean_ctor_get(x_140, 0); +lean_inc(x_146); +lean_dec(x_140); +x_1 = x_146; +x_7 = x_145; +goto _start; +} +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_148 = lean_ctor_get(x_139, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_139, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_150 = x_139; +} else { + lean_dec_ref(x_139); + 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_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_10); +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_97); +return x_153; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___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_Lean_Compiler_visitLetImp_go___at_Lean_Compiler_Simp_inlineProjInst_x3f_visit___spec__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_Compiler_Simp_inlineProjInst_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: +{ +if (lean_obj_tag(x_1) == 11) +{ +lean_object* x_8; 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_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = lean_st_ref_get(x_6, x_7); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_st_ref_get(x_4, 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); +lean_inc(x_6); +lean_inc(x_5); +x_15 = l_Lean_Compiler_InferType_inferType(x_8, x_14, x_5, x_6, x_13); +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_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Compiler_isClass_x3f(x_16, x_5, x_6, x_17); +lean_dec(x_16); +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_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_dec(x_19); +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); +lean_dec(x_18); +x_27 = lean_st_ref_get(x_6, x_26); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_st_ref_get(x_4, x_28); +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 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_33 = l_Lean_Compiler_InferType_inferType(x_1, x_32, x_5, x_6, x_31); +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_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Compiler_isClass_x3f(x_34, x_5, x_6, x_35); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; 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_104; lean_object* x_105; 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_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_80 = lean_st_ref_get(x_6, x_38); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_st_ref_get(x_4, x_81); +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_123 = lean_st_ref_get(x_6, x_84); +x_124 = lean_ctor_get(x_123, 1); +lean_inc(x_124); +lean_dec(x_123); +x_125 = lean_st_ref_take(x_4, x_124); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_128 = lean_ctor_get(x_126, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_126, 2); +lean_inc(x_129); +lean_dec(x_126); +x_130 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_131 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = lean_st_ref_set(x_4, x_131, x_127); +x_133 = lean_ctor_get(x_132, 1); +lean_inc(x_133); +lean_dec(x_132); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_134 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj(x_1, x_2, x_3, x_4, x_5, x_6, x_133); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_135; +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = lean_box(0); +x_85 = x_137; +x_86 = x_136; +goto block_103; +} +else +{ +lean_object* x_138; uint8_t x_139; +x_138 = lean_ctor_get(x_134, 1); +lean_inc(x_138); +lean_dec(x_134); +x_139 = !lean_is_exclusive(x_135); +if (x_139 == 0) +{ +lean_object* x_140; lean_object* x_141; +x_140 = lean_ctor_get(x_135, 0); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_141 = l_Lean_Compiler_mkLetUsingScope(x_140, x_4, x_5, x_6, x_138); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +lean_ctor_set(x_135, 0, x_142); +x_85 = x_135; +x_86 = x_143; +goto block_103; +} +else +{ +lean_object* x_144; lean_object* x_145; +lean_free_object(x_135); +x_144 = lean_ctor_get(x_141, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_141, 1); +lean_inc(x_145); +lean_dec(x_141); +x_104 = x_144; +x_105 = x_145; +goto block_122; +} +} +else +{ +lean_object* x_146; lean_object* x_147; +x_146 = lean_ctor_get(x_135, 0); +lean_inc(x_146); +lean_dec(x_135); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_147 = l_Lean_Compiler_mkLetUsingScope(x_146, x_4, x_5, x_6, x_138); +if (lean_obj_tag(x_147) == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_150, 0, x_148); +x_85 = x_150; +x_86 = x_149; +goto block_103; +} +else +{ +lean_object* x_151; lean_object* x_152; +x_151 = lean_ctor_get(x_147, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_147, 1); +lean_inc(x_152); +lean_dec(x_147); +x_104 = x_151; +x_105 = x_152; +goto block_122; +} +} +} +} +else +{ +lean_object* x_153; lean_object* x_154; +x_153 = lean_ctor_get(x_134, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_134, 1); +lean_inc(x_154); +lean_dec(x_134); +x_104 = x_153; +x_105 = x_154; +goto block_122; +} +block_79: +{ +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 0); +lean_dec(x_42); +x_43 = lean_box(0); +lean_ctor_set(x_39, 0, x_43); +return x_39; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +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_44); +return x_46; +} +} +else +{ +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_48 = !lean_is_exclusive(x_40); +if (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_40, 0); +x_50 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_47); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = l_Lean_Compiler_Simp_internalize_visitLambda(x_49, x_2, x_3, x_4, x_5, x_6, x_51); +if (lean_obj_tag(x_52) == 0) +{ +uint8_t x_53; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; +x_54 = lean_ctor_get(x_52, 0); +lean_ctor_set(x_40, 0, x_54); +lean_ctor_set(x_52, 0, x_40); +return x_52; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_52, 0); +x_56 = lean_ctor_get(x_52, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_52); +lean_ctor_set(x_40, 0, x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_40); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +else +{ +uint8_t x_58; +lean_free_object(x_40); +x_58 = !lean_is_exclusive(x_52); +if (x_58 == 0) +{ +return x_52; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_52, 0); +x_60 = lean_ctor_get(x_52, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_52); +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 +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_40, 0); +lean_inc(x_62); +lean_dec(x_40); +x_63 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_47); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Lean_Compiler_Simp_internalize_visitLambda(x_62, x_2, x_3, x_4, x_5, x_6, x_64); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_68 = x_65; +} else { + lean_dec_ref(x_65); + x_68 = lean_box(0); +} +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_66); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_68; +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = lean_ctor_get(x_65, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_65, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + 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, 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; +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_75 = !lean_is_exclusive(x_39); +if (x_75 == 0) +{ +return x_39; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_39, 0); +x_77 = lean_ctor_get(x_39, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_39); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +block_103: +{ +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; uint8_t x_99; +x_87 = lean_st_ref_get(x_6, x_86); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_get(x_4, x_88); +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_ctor_get(x_83, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_83, 1); +lean_inc(x_93); +lean_dec(x_83); +x_94 = lean_ctor_get(x_90, 2); +lean_inc(x_94); +lean_dec(x_90); +x_95 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +lean_ctor_set(x_95, 2, x_94); +x_96 = lean_st_ref_get(x_6, x_91); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = lean_st_ref_set(x_4, x_95, x_97); +x_99 = !lean_is_exclusive(x_98); +if (x_99 == 0) +{ +lean_object* x_100; +x_100 = lean_ctor_get(x_98, 0); +lean_dec(x_100); +lean_ctor_set(x_98, 0, x_85); +x_39 = x_98; +goto block_79; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_98, 1); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_85); +lean_ctor_set(x_102, 1, x_101); +x_39 = x_102; +goto block_79; +} +} +block_122: +{ +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; uint8_t x_118; +x_106 = lean_st_ref_get(x_6, x_105); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +lean_dec(x_106); +x_108 = lean_st_ref_get(x_4, x_107); +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_83, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_83, 1); +lean_inc(x_112); +lean_dec(x_83); +x_113 = lean_ctor_get(x_109, 2); +lean_inc(x_113); +lean_dec(x_109); +x_114 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_112); +lean_ctor_set(x_114, 2, x_113); +x_115 = lean_st_ref_get(x_6, x_110); +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +x_117 = lean_st_ref_set(x_4, x_114, x_116); +x_118 = !lean_is_exclusive(x_117); +if (x_118 == 0) +{ +lean_object* x_119; +x_119 = lean_ctor_get(x_117, 0); +lean_dec(x_119); +lean_ctor_set_tag(x_117, 1); +lean_ctor_set(x_117, 0, x_104); +x_39 = x_117; +goto block_79; +} +else +{ +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_117, 1); +lean_inc(x_120); +lean_dec(x_117); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_104); +lean_ctor_set(x_121, 1, x_120); +x_39 = x_121; +goto block_79; +} +} +} +else +{ +uint8_t x_155; +lean_dec(x_37); +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_155 = !lean_is_exclusive(x_36); +if (x_155 == 0) +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_36, 0); +lean_dec(x_156); +x_157 = lean_box(0); +lean_ctor_set(x_36, 0, x_157); +return x_36; +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_36, 1); +lean_inc(x_158); +lean_dec(x_36); +x_159 = lean_box(0); +x_160 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +return x_160; +} +} +} +else +{ +uint8_t x_161; +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_161 = !lean_is_exclusive(x_33); +if (x_161 == 0) +{ +return x_33; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_33, 0); +x_163 = lean_ctor_get(x_33, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_33); +x_164 = lean_alloc_ctor(1, 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; +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_165 = !lean_is_exclusive(x_15); +if (x_165 == 0) +{ +return x_15; +} +else +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_15, 0); +x_167 = lean_ctor_get(x_15, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_15); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +return x_168; +} +} +} +else +{ +lean_object* x_169; lean_object* x_170; +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_169 = lean_box(0); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_7); +return x_170; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_betaReduce(lean_object* x_1, lean_object* x_2, 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_beta(x_1, x_2); +x_10 = l_Lean_Compiler_Simp_internalize_visitLambda(x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_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, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_4, 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; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_local_ctx_find(x_13, x_1); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_10, 0, x_15); +return x_10; +} +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; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_10); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_local_ctx_find(x_18, x_1); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, 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_17); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_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: +{ +switch (lean_obj_tag(x_1)) { +case 6: +{ +lean_object* x_8; +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_dec(x_1); +x_1 = x_8; +goto _start; +} +case 8: +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +switch (lean_obj_tag(x_10)) { +case 5: +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 1) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +if (lean_obj_tag(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, 2); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 0); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_14); +lean_dec(x_13); +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_7); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_14); +x_20 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec(x_21); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; +lean_dec(x_14); +lean_dec(x_13); +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_20, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_20, 0, x_25); +return x_20; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +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_29; +x_29 = !lean_is_exclusive(x_20); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_20, 1); +x_31 = lean_ctor_get(x_20, 0); +lean_dec(x_31); +x_32 = !lean_is_exclusive(x_22); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_22, 0); +x_34 = l_Lean_LocalDecl_isJp(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; +lean_free_object(x_22); +lean_dec(x_14); +lean_dec(x_13); +x_35 = lean_box(0); +lean_ctor_set(x_20, 0, x_35); +return x_20; +} +else +{ +lean_object* x_36; uint8_t x_37; +lean_free_object(x_20); +x_36 = lean_st_ref_get(x_6, x_30); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +lean_dec(x_38); +x_40 = l_Lean_Expr_isConstructorApp(x_39, x_13); +lean_dec(x_13); +if (x_40 == 0) +{ +lean_object* x_41; +lean_free_object(x_22); +lean_dec(x_14); +x_41 = lean_box(0); +lean_ctor_set(x_36, 0, x_41); +return x_36; +} +else +{ +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_36, 0, x_22); +return x_36; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_42 = lean_ctor_get(x_36, 0); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_36); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_Expr_isConstructorApp(x_44, x_13); +lean_dec(x_13); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +lean_free_object(x_22); +lean_dec(x_14); +x_46 = lean_box(0); +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; +lean_ctor_set(x_22, 0, x_14); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_22); +lean_ctor_set(x_48, 1, x_43); +return x_48; +} +} +} +} +else +{ +lean_object* x_49; uint8_t x_50; +x_49 = lean_ctor_get(x_22, 0); +lean_inc(x_49); +lean_dec(x_22); +x_50 = l_Lean_LocalDecl_isJp(x_49); +lean_dec(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +lean_dec(x_14); +lean_dec(x_13); +x_51 = lean_box(0); +lean_ctor_set(x_20, 0, x_51); +return x_20; +} +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; +lean_free_object(x_20); +x_52 = lean_st_ref_get(x_6, x_30); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_55 = x_52; +} else { + lean_dec_ref(x_52); + x_55 = lean_box(0); +} +x_56 = lean_ctor_get(x_53, 0); +lean_inc(x_56); +lean_dec(x_53); +x_57 = l_Lean_Expr_isConstructorApp(x_56, x_13); +lean_dec(x_13); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_14); +x_58 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_59 = lean_alloc_ctor(0, 2, 0); +} else { + x_59 = x_55; +} +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_14); +if (lean_is_scalar(x_55)) { + x_61 = lean_alloc_ctor(0, 2, 0); +} else { + x_61 = x_55; +} +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_54); +return x_61; +} +} +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_62 = lean_ctor_get(x_20, 1); +lean_inc(x_62); +lean_dec(x_20); +x_63 = lean_ctor_get(x_22, 0); +lean_inc(x_63); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + x_64 = x_22; +} else { + lean_dec_ref(x_22); + x_64 = lean_box(0); +} +x_65 = l_Lean_LocalDecl_isJp(x_63); +lean_dec(x_63); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_64); +lean_dec(x_14); +lean_dec(x_13); +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_62); +return x_67; +} +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_st_ref_get(x_6, x_62); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_71 = x_68; +} else { + lean_dec_ref(x_68); + x_71 = lean_box(0); +} +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +lean_dec(x_69); +x_73 = l_Lean_Expr_isConstructorApp(x_72, x_13); +lean_dec(x_13); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; +lean_dec(x_64); +lean_dec(x_14); +x_74 = lean_box(0); +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; +} +else +{ +lean_object* x_76; lean_object* x_77; +if (lean_is_scalar(x_64)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_64; +} +lean_ctor_set(x_76, 0, x_14); +if (lean_is_scalar(x_71)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_71; +} +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_70); +return x_77; +} +} +} +} +} +} +else +{ +lean_object* x_78; lean_object* x_79; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_1); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_7); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_7); +return x_81; +} +} +case 8: +{ +lean_dec(x_1); +x_1 = x_10; +goto _start; +} +default: +{ +lean_object* x_83; lean_object* x_84; +lean_dec(x_10); +lean_dec(x_1); +x_83 = lean_box(0); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_7); +return x_84; +} +} +} +default: +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_1); +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_7); +return x_86; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__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_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f(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_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 6: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +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_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +x_7 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(x_5, x_2); +x_8 = l_Lean_Expr_lam___override(x_3, x_4, x_7, x_6); +return x_8; +} +case 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_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_ctor_get(x_1, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_12); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_dec(x_1); +x_14 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(x_12, x_2); +x_15 = l_Lean_Expr_letE___override(x_9, x_10, x_11, x_14, x_13); +return x_15; +} +default: +{ +lean_dec(x_1); +lean_inc(x_2); +return x_2; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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, 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_nat_dec_le(x_4, x_3); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_2, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_27; lean_object* x_91; uint8_t x_92; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_2, x_16); +lean_dec(x_2); +x_91 = lean_array_get_size(x_1); +x_92 = lean_nat_dec_lt(x_3, x_91); +lean_dec(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; +x_93 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_94 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_93); +x_27 = x_94; +goto block_90; +} +else +{ +lean_object* x_95; +x_95 = lean_array_fget(x_1, x_3); +x_27 = x_95; +goto block_90; +} +block_26: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +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_19); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_17; +x_3 = x_24; +x_6 = x_23; +x_12 = x_19; +goto _start; +} +} +block_90: +{ +lean_object* x_28; lean_object* x_29; +x_28 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f(x_27, x_7, x_8, x_9, x_10, x_11, x_12); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_6); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_box(0); +x_18 = x_31; +x_19 = x_30; +goto block_26; +} +else +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_32 = lean_ctor_get(x_28, 1); +lean_inc(x_32); +lean_dec(x_28); +x_33 = lean_ctor_get(x_29, 0); +lean_inc(x_33); +lean_dec(x_29); +lean_inc(x_33); +x_34 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__1(x_33, x_7, x_8, x_9, x_10, x_11, x_32); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_box(0); +x_18 = x_38; +x_19 = x_37; +goto block_26; } 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; -x_12 = l_Lean_LocalDecl_userName(x_1); -x_13 = lean_st_ref_get(x_6, x_7); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_st_ref_get(x_3, x_14); -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_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -lean_inc(x_12); -x_19 = l_Std_HashMapImp_find_x3f___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goTailApp___spec__1(x_18, x_12); -if (lean_obj_tag(x_19) == 0) +lean_object* x_39; uint8_t x_40; +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); +lean_dec(x_34); +x_40 = !lean_is_exclusive(x_36); +if (x_40 == 0) { -lean_object* x_20; -lean_inc(x_6); -x_20 = l_Lean_Compiler_getLCNFSize(x_8, x_5, x_6, x_17); -if (lean_obj_tag(x_20) == 0) +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_36, 0); +x_42 = l_Lean_LocalDecl_value(x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 6) { -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; -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_6, x_22); -lean_dec(x_6); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_st_ref_take(x_3, x_24); -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_is_exclusive(x_26); -if (x_28 == 0) +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_42, 2); +lean_inc(x_43); +lean_dec(x_42); +lean_inc(x_11); +lean_inc(x_10); +x_44 = l_Lean_Compiler_isCasesApp_x3f(x_43, x_10, x_11, x_39); +if (lean_obj_tag(x_44) == 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_26, 0); -x_30 = lean_ctor_get(x_26, 1); -x_31 = lean_unsigned_to_nat(1u); -lean_inc(x_12); -x_32 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_29, x_12, x_31); -x_33 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_30, x_12, x_21); -lean_ctor_set(x_26, 1, x_33); -lean_ctor_set(x_26, 0, x_32); -x_34 = lean_st_ref_set(x_3, x_26, x_27); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_object* x_45; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -x_37 = lean_box(0); -lean_ctor_set(x_34, 0, x_37); -return x_34; +lean_object* x_46; lean_object* x_47; +lean_free_object(x_36); +lean_dec(x_33); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_box(0); +x_18 = x_47; +x_19 = x_46; +goto block_26; } 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 = lean_box(0); -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; -} +uint8_t x_48; +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +lean_dec(x_49); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_dec(x_44); +lean_ctor_set(x_45, 0, x_33); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_45); +lean_ctor_set(x_36, 0, x_51); +x_18 = x_36; +x_19 = x_50; +goto block_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; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_41 = lean_ctor_get(x_26, 0); -x_42 = lean_ctor_get(x_26, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_26); -x_43 = lean_unsigned_to_nat(1u); -lean_inc(x_12); -x_44 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_41, x_12, x_43); -x_45 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_42, x_12, x_21); -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_st_ref_set(x_3, x_46, x_27); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_49 = x_47; -} else { - lean_dec_ref(x_47); - x_49 = lean_box(0); -} -x_50 = lean_box(0); -if (lean_is_scalar(x_49)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_49; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_45); +x_52 = lean_ctor_get(x_44, 1); +lean_inc(x_52); +lean_dec(x_44); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_33); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_36, 0, x_54); +x_18 = x_36; +x_19 = x_52; +goto block_26; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; } } else { -uint8_t x_52; -lean_dec(x_12); -lean_dec(x_6); -x_52 = !lean_is_exclusive(x_20); -if (x_52 == 0) +uint8_t x_55; +lean_free_object(x_36); +lean_dec(x_33); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +x_55 = !lean_is_exclusive(x_44); +if (x_55 == 0) { -return x_20; +return x_44; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_20, 0); -x_54 = lean_ctor_get(x_20, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_20); -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_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_44, 0); +x_57 = lean_ctor_get(x_44, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_44); +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; } } } 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_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -lean_dec(x_8); -lean_dec(x_5); -x_56 = lean_ctor_get(x_19, 0); -lean_inc(x_56); -lean_dec(x_19); -x_57 = lean_st_ref_get(x_6, x_17); -lean_dec(x_6); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_st_ref_take(x_3, x_58); -x_60 = lean_ctor_get(x_59, 0); +lean_object* x_59; +lean_dec(x_42); +lean_free_object(x_36); +lean_dec(x_33); +x_59 = lean_box(0); +x_18 = x_59; +x_19 = x_39; +goto block_26; +} +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_36, 0); lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_60, 0); -lean_inc(x_62); -x_63 = lean_unsigned_to_nat(1u); -x_64 = lean_nat_add(x_56, x_63); -lean_dec(x_56); -x_65 = l_Std_HashMap_insert___at_Lean_Compiler_JoinPoints_JoinPointFinder_findJoinPoints_goLetValue___spec__46(x_62, x_12, x_64); -x_66 = lean_ctor_get(x_60, 1); -lean_inc(x_66); +lean_dec(x_36); +x_61 = l_Lean_LocalDecl_value(x_60); lean_dec(x_60); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -x_68 = lean_st_ref_set(x_3, x_67, x_61); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +if (lean_obj_tag(x_61) == 6) { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_68, 0); -lean_dec(x_70); -x_71 = lean_box(0); -lean_ctor_set(x_68, 0, x_71); -return x_68; +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_61, 2); +lean_inc(x_62); +lean_dec(x_61); +lean_inc(x_11); +lean_inc(x_10); +x_63 = l_Lean_Compiler_isCasesApp_x3f(x_62, x_10, x_11, x_39); +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; lean_object* x_66; +lean_dec(x_33); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_box(0); +x_18 = x_66; +x_19 = x_65; +goto block_26; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_68, 1); -lean_inc(x_72); -lean_dec(x_68); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_72); -return x_74; -} +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + x_67 = x_64; +} else { + lean_dec_ref(x_64); + x_67 = lean_box(0); } +x_68 = lean_ctor_get(x_63, 1); +lean_inc(x_68); +lean_dec(x_63); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(1, 1, 0); +} else { + x_69 = x_67; } +lean_ctor_set(x_69, 0, x_33); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_69); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_18 = x_71; +x_19 = x_68; +goto block_26; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Meta", 4); -return x_1; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_33); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +x_72 = lean_ctor_get(x_63, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_63, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_74 = x_63; +} else { + lean_dec_ref(x_63); + x_74 = lean_box(0); } +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(1, 2, 0); +} else { + x_75 = x_74; } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_ctor_set(x_75, 0, x_72); +lean_ctor_set(x_75, 1, x_73); +return x_75; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("debug", 5); -return x_1; +lean_object* x_76; +lean_dec(x_61); +lean_dec(x_33); +x_76 = lean_box(0); +x_18 = x_76; +x_19 = x_39; +goto block_26; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2; -x_2 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("found decl ", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6() { -_start: +lean_object* x_77; uint8_t x_78; +x_77 = lean_ctor_get(x_28, 1); +lean_inc(x_77); +lean_dec(x_28); +x_78 = !lean_is_exclusive(x_29); +if (x_78 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_79 = lean_ctor_get(x_29, 0); +x_80 = lean_ctor_get(x_6, 0); +lean_inc(x_80); +x_81 = lean_name_eq(x_79, x_80); +lean_dec(x_80); +lean_dec(x_79); +if (x_81 == 0) +{ +lean_object* x_82; +lean_free_object(x_29); +lean_dec(x_6); +x_82 = lean_box(0); +x_18 = x_82; +x_19 = x_77; +goto block_26; } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_83; +x_83 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_83, 0, x_6); +lean_ctor_set(x_29, 0, x_83); +x_18 = x_29; +x_19 = x_77; +goto block_26; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue(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: -{ -switch (lean_obj_tag(x_1)) { -case 5: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_Lean_Expr_getAppFn(x_1); -lean_dec(x_1); -x_8 = l_Lean_Compiler_Simp_findLambda_x3f(x_7, x_3, x_4, x_5, x_6); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -if (lean_obj_tag(x_9) == 0) +else { -uint8_t x_10; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = !lean_is_exclusive(x_8); -if (x_10 == 0) +lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_84 = lean_ctor_get(x_29, 0); +lean_inc(x_84); +lean_dec(x_29); +x_85 = lean_ctor_get(x_6, 0); +lean_inc(x_85); +x_86 = lean_name_eq(x_84, x_85); +lean_dec(x_85); +lean_dec(x_84); +if (x_86 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_8, 0); -lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_8, 0, x_12); -return x_8; +lean_object* x_87; +lean_dec(x_6); +x_87 = lean_box(0); +x_18 = x_87; +x_19 = x_77; +goto block_26; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -lean_dec(x_8); -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_13); -return x_15; +lean_object* x_88; lean_object* x_89; +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_6); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, x_88); +x_18 = x_89; +x_19 = x_77; +goto block_26; +} +} +} +} } } else { -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_16 = lean_ctor_get(x_8, 1); -lean_inc(x_16); -lean_dec(x_8); -x_17 = lean_ctor_get(x_9, 0); -lean_inc(x_17); -lean_dec(x_9); -x_18 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4; -x_19 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1(x_18, x_2, x_3, x_4, x_5, x_16); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -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_Compiler_Simp_collectInlineStats_goValue___lambda__1(x_17, x_23, x_2, x_3, x_4, x_5, x_22); +lean_object* x_96; lean_object* x_97; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_17); -return x_24; +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_6); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_12); +return x_97; +} } 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_35; -x_25 = lean_ctor_get(x_19, 1); -lean_inc(x_25); -lean_dec(x_19); -x_26 = l_Lean_LocalDecl_userName(x_17); -x_27 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6; -x_29 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -x_30 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7; -x_31 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2(x_18, x_31, x_2, x_3, x_4, x_5, x_25); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Compiler_Simp_collectInlineStats_goValue___lambda__1(x_17, x_33, x_2, x_3, x_4, x_5, x_34); +lean_object* x_98; lean_object* x_99; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_33); -lean_dec(x_17); -return x_35; +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_6); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_12); +return x_99; } } } -case 6: +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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, 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_36; -x_36 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_1, x_2, x_3, x_4, x_5, x_6); -return x_36; -} -default: +uint8_t x_13; +x_13 = lean_nat_dec_le(x_4, x_3); +if (x_13 == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_2, 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_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_2, x_16); lean_dec(x_2); -lean_dec(x_1); -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); -return x_38; -} +x_18 = lean_array_get_size(x_6); +x_19 = lean_nat_dec_lt(x_3, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_17; +x_3 = x_20; +goto _start; } +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; +x_22 = lean_array_fget(x_6, x_3); +x_23 = lean_box(0); +x_24 = lean_array_fset(x_6, x_3, x_23); +x_25 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f_inlineJp(x_22, x_1); +x_26 = lean_array_fset(x_24, x_3, x_25); +x_27 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_17; +x_3 = x_27; +x_6 = x_26; +goto _start; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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 { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +lean_object* x_29; lean_object* x_30; +lean_dec(x_3); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_6); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_12); +return x_30; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +lean_object* x_31; lean_object* x_32; +lean_dec(x_3); +lean_dec(x_2); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_6); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_12); +return x_32; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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) { +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_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, lean_object* x_9) { _start: { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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_5); -lean_dec(x_4); +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_box(0); +x_11 = lean_ctor_get(x_1, 3); +lean_inc(x_11); lean_dec(x_1); -return x_12; -} +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 2); +lean_inc(x_14); +lean_dec(x_11); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_13); +lean_inc(x_12); +x_15 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__1(x_3, x_12, x_13, x_12, x_14, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +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_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +lean_ctor_set(x_15, 0, x_10); +return x_15; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__5(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; -} +lean_object* x_21; +x_21 = lean_ctor_get(x_16, 0); +lean_inc(x_21); +lean_dec(x_16); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_15, 0); +lean_dec(x_23); +lean_ctor_set(x_15, 0, x_10); +return x_15; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_dec(x_15); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__7(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_15, 1); +lean_inc(x_26); +lean_dec(x_15); +x_27 = lean_ctor_get(x_21, 0); +lean_inc(x_27); +lean_dec(x_21); +x_28 = l_Lean_Compiler_findDecl_x3f___at_Lean_Compiler_Simp_simpCasesOnCases_x3f_isJpCtor_x3f___spec__1(x_27, x_4, x_5, x_6, x_7, x_8, x_26); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_28, 0); +lean_dec(x_32); +lean_ctor_set(x_28, 0, x_10); +return x_28; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__8(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_10); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__9(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; -} +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (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_28, 1); +x_37 = lean_ctor_get(x_28, 0); +lean_dec(x_37); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +x_39 = l_Lean_LocalDecl_value(x_38); +lean_dec(x_38); +if (lean_obj_tag(x_39) == 6) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +lean_free_object(x_28); +x_40 = lean_ctor_get(x_39, 2); +lean_inc(x_40); +lean_dec(x_39); +lean_inc(x_12); +x_41 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__2(x_40, x_12, x_13, x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_40); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_41, 0); +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, 0); +x_46 = l_Lean_mkAppN(x_2, x_45); +lean_ctor_set(x_43, 0, x_46); +return x_41; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__10(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +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_Lean_mkAppN(x_2, x_47); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_41, 0, x_49); +return x_41; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___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, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__11(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_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +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_50 = lean_ctor_get(x_41, 0); +x_51 = lean_ctor_get(x_41, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_41); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + x_53 = x_50; +} else { + lean_dec_ref(x_50); + x_53 = lean_box(0); } +x_54 = l_Lean_mkAppN(x_2, x_52); +if (lean_is_scalar(x_53)) { + x_55 = lean_alloc_ctor(1, 1, 0); +} else { + x_55 = x_53; } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_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) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_Compiler_Simp_collectInlineStats_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; +lean_ctor_set(x_55, 0, 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_51); +return x_56; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___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: +else { -lean_object* x_7; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_39); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); -return x_7; +lean_ctor_set(x_28, 0, x_10); +return x_28; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___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: +else { -lean_object* x_8; -x_8 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2(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); -return x_8; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_28, 1); +lean_inc(x_57); +lean_dec(x_28); +x_58 = lean_ctor_get(x_30, 0); +lean_inc(x_58); +lean_dec(x_30); +x_59 = l_Lean_LocalDecl_value(x_58); +lean_dec(x_58); +if (lean_obj_tag(x_59) == 6) +{ +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; +x_60 = lean_ctor_get(x_59, 2); +lean_inc(x_60); +lean_dec(x_59); +lean_inc(x_12); +x_61 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__2(x_60, x_12, x_13, x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_60); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_64 = x_61; +} else { + lean_dec_ref(x_61); + x_64 = lean_box(0); } +x_65 = lean_ctor_get(x_62, 0); +lean_inc(x_65); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + x_66 = x_62; +} else { + lean_dec_ref(x_62); + x_66 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats_goValue___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: +x_67 = l_Lean_mkAppN(x_2, x_65); +if (lean_is_scalar(x_66)) { + x_68 = lean_alloc_ctor(1, 1, 0); +} else { + x_68 = x_66; +} +lean_ctor_set(x_68, 0, x_67); +if (lean_is_scalar(x_64)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_64; +} +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_63); +return x_69; +} +else { -lean_object* x_8; -x_8 = l_Lean_Compiler_Simp_collectInlineStats_goValue___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); +lean_object* x_70; +lean_dec(x_59); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_8; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_10); +lean_ctor_set(x_70, 1, x_57); +return x_70; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_Simp_InlineStats_numOccs___default___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_Lean_Compiler_Simp_collectInlineStats___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2; -x_2 = lean_unsigned_to_nat(0u); -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_Compiler_Simp_collectInlineStats___closed__3() { -_start: +else { -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; -} +uint8_t x_71; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_71 = !lean_is_exclusive(x_15); +if (x_71 == 0) +{ +return x_15; } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_Simp_collectInlineStats___closed__3; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_15, 0); +x_73 = lean_ctor_get(x_15, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_15); +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; +} +} } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats___closed__5() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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, lean_object* x_11, lean_object* x_12) { _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_Compiler_Simp_collectInlineStats___closed__4; -x_3 = l_Lean_Compiler_Simp_collectInlineStats___closed__3; -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; +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__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); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats___closed__6() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_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, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_collectInlineStats___closed__2; -x_2 = l_Lean_Compiler_Simp_collectInlineStats___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; +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_simpCasesOnCases_x3f___spec__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_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Lean_Compiler_Simp_collectInlineStats___closed__7() { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpCasesOnCases_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: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Compiler_Simp_collectInlineStats___closed__6; -x_2 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_3 = lean_unsigned_to_nat(1u); -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_object* x_10; +x_10 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_collectInlineStats(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_simpValue_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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_5 = lean_st_ref_get(x_3, x_4); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Compiler_Simp_collectInlineStats___closed__7; -x_8 = lean_st_mk_ref(x_7, x_6); +lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_8 = l_Lean_Compiler_Simp_simpProj_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_st_ref_get(x_3, x_10); -x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_1); +x_11 = l_Lean_Compiler_Simp_simpAppApp_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_10); +x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Compiler_Simp_collectInlineStats___closed__1; -x_14 = lean_st_mk_ref(x_13, x_12); -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_3); -lean_inc(x_9); -lean_inc(x_15); -x_17 = l_Lean_Compiler_Simp_collectInlineStats_goLambda(x_1, x_15, x_9, x_2, x_3, 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_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_get(x_3, x_18); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_st_ref_get(x_15, x_20); -lean_dec(x_15); -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_st_ref_get(x_3, x_23); -lean_dec(x_3); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_st_ref_get(x_9, x_25); -lean_dec(x_9); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -lean_ctor_set(x_26, 0, x_22); -return x_26; -} -else +if (lean_obj_tag(x_12) == 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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_22); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Compiler_Simp_inlineProjInst_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_13); +return x_14; } else { -uint8_t x_31; -lean_dec(x_15); -lean_dec(x_9); +uint8_t x_15; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_31 = !lean_is_exclusive(x_17); -if (x_31 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_11, 0); +lean_dec(x_16); +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) { -return x_17; +return x_11; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_17, 0); -x_33 = lean_ctor_get(x_17, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_17); -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_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_12, 0); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_11, 0, x_19); +return x_11; } } -static lean_object* _init_l_Lean_Compiler_Simp_Config_increaseFactor___default() { -_start: +else { -lean_object* x_1; -x_1 = lean_unsigned_to_nat(2u); -return x_1; +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_11, 1); +lean_inc(x_20); +lean_dec(x_11); +x_21 = lean_ctor_get(x_12, 0); +lean_inc(x_21); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + x_22 = x_12; +} else { + lean_dec_ref(x_12); + x_22 = lean_box(0); } +if (lean_is_scalar(x_22)) { + x_23 = lean_alloc_ctor(1, 1, 0); +} else { + x_23 = x_22; } -static lean_object* _init_l_Lean_Compiler_Simp_Context_config___default() { -_start: -{ -lean_object* x_1; -x_1 = lean_unsigned_to_nat(2u); -return x_1; +lean_ctor_set(x_23, 0, 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_20); +return x_24; } } -static uint8_t _init_l_Lean_Compiler_Simp_Context_localInline___default() { -_start: -{ -uint8_t x_1; -x_1 = 1; -return x_1; } +else +{ +uint8_t x_25; +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_8); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_8, 0); +lean_dec(x_26); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +return x_8; } -static uint8_t _init_l_Lean_Compiler_Simp_State_simplified___default() { -_start: +else { -uint8_t x_1; -x_1 = 0; -return x_1; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_9, 0); +lean_inc(x_28); +lean_dec(x_9); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_8, 0, x_29); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -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; uint8_t x_13; -x_6 = lean_st_ref_get(x_4, x_5); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_st_ref_take(x_1, x_7); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); +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_8, 1); +lean_inc(x_30); lean_dec(x_8); -x_10 = 1; -x_11 = lean_box(x_10); -x_12 = lean_st_ref_set(x_1, x_11, x_9); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_12, 0, x_15); -return x_12; +x_31 = lean_ctor_get(x_9, 0); +lean_inc(x_31); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + x_32 = x_9; +} else { + lean_dec_ref(x_9); + x_32 = lean_box(0); } -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -lean_dec(x_12); -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; +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(1, 1, 0); +} else { + x_33 = x_32; } +lean_ctor_set(x_33, 0, x_31); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +return x_34; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_markSimplified___rarg___boxed), 5, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__1() { _start: { -lean_object* x_6; -x_6 = l_Lean_Compiler_Simp_markSimplified___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_6; +lean_object* x_1; +x_1 = l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_markSimplified___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2() { _start: { -lean_object* x_2; -x_2 = l_Lean_Compiler_Simp_markSimplified(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___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_Lean_Compiler_Simp_shouldInline(lean_object* x_1, 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_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3() { _start: { -uint8_t x_8; -x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -if (x_8 == 0) -{ -uint8_t x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_2); -x_9 = 0; -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_7); -return x_11; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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 -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_2, 1); -lean_inc(x_12); -lean_dec(x_2); -x_13 = l_Lean_LocalDecl_userName(x_1); -x_14 = l_Lean_Compiler_Simp_InlineStats_shouldInline(x_12, x_13); -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; } +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_Compiler_Simp_shouldInline___boxed(lean_object* x_1, 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_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5() { _start: { -lean_object* x_8; -x_8 = l_Lean_Compiler_Simp_shouldInline(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_1); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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_Compiler_Simp_inlineCandidate_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6() { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_10); -x_12 = lean_nat_sub(x_11, x_2); -lean_dec(x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, 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_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3; +x_3 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4; +x_4 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5; +x_5 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set(x_5, 4, x_4); +lean_ctor_set(x_5, 5, x_2); +lean_ctor_set(x_5, 6, x_3); +lean_ctor_set(x_5, 7, x_3); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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_throwError___at_Lean_Compiler_Simp_visitLambda___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_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_9); -x_11 = lean_nat_dec_lt(x_10, x_2); +lean_object* x_8; 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_8 = lean_ctor_get(x_5, 5); +x_9 = lean_st_ref_get(x_6, 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_10, 0); +lean_inc(x_12); lean_dec(x_10); -if (x_11 == 0) +x_13 = lean_st_ref_get(x_6, x_11); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_get(x_4, x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__1(x_1, x_2, x_12, 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_13; +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_15, 0); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_ctor_get(x_5, 2); +x_20 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6; +lean_inc(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_20); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_19); +x_22 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_1); +lean_inc(x_8); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_23); +return x_15; } else { -lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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_8); -return x_15; +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_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_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_5, 2); +x_28 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6; +lean_inc(x_27); +x_29 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_29, 0, x_12); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_27); +x_30 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_1); +lean_inc(x_8); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_8); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_25); +return x_32; +} } } +static lean_object* _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("type expected", 13); +return x_1; +} } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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) { +static lean_object* _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2() { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_LocalDecl_value(x_1); -x_11 = l_Lean_Compiler_getLambdaArity(x_10); -lean_dec(x_10); -x_12 = lean_apply_7(x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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) { +static lean_object* _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3() { _start: { -lean_object* x_10; -lean_inc(x_8); -lean_inc(x_7); -x_10 = l_Lean_Compiler_getStage1Decl_x3f(x_1, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___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_11; +lean_object* x_8; 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_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_4, x_9); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_14 = l_Lean_Compiler_InferType_inferType(x_1, x_13, x_5, x_6, x_12); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 3) +{ +uint8_t x_16; lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -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_dec(x_1); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) { -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; -} +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 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_14, 0, x_18); +return x_14; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_ctor_get(x_11, 0); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 1); lean_inc(x_19); -lean_dec(x_11); -x_20 = l_Lean_Compiler_Decl_getArity(x_19); -lean_dec(x_19); -x_21 = lean_apply_7(x_2, x_20, x_4, x_5, x_6, x_7, x_8, x_18); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +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; } } else { uint8_t x_22; -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_22 = !lean_is_exclusive(x_10); +x_22 = !lean_is_exclusive(x_14); if (x_22 == 0) { -return x_10; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_10, 0); -x_24 = lean_ctor_get(x_10, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_10); -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_Lean_Compiler_Simp_inlineCandidate_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_8; lean_object* x_9; -x_8 = l_Lean_Expr_getAppFn(x_1); -x_9 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__2___boxed), 8, 1); -lean_closure_set(x_9, 0, x_1); -if (lean_obj_tag(x_8) == 4) -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_st_ref_get(x_6, x_7); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; -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); -x_16 = 0; -lean_inc(x_10); -x_17 = l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(x_15, x_16, x_10); -if (x_17 == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_14, 1); +x_24 = lean_ctor_get(x_14, 0); +lean_dec(x_24); +x_25 = l_Lean_Expr_isAnyType(x_15); +lean_dec(x_15); +if (x_25 == 0) { -lean_object* x_18; -lean_dec(x_10); -lean_dec(x_9); +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_free_object(x_14); +x_26 = l_Lean_indentExpr(x_1); +x_27 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2; +x_28 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_30 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3(x_30, x_2, x_3, x_4, x_5, x_6, x_23); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_18 = lean_box(0); -lean_ctor_set(x_11, 0, x_18); -return x_11; +return x_31; } else { -lean_object* x_19; lean_object* x_20; -lean_free_object(x_11); -x_19 = lean_box(0); -x_20 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__4(x_10, x_9, x_19, x_2, x_3, x_4, x_5, x_6, x_14); -return x_20; +lean_object* x_32; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_32 = l_Lean_levelOne; +lean_ctor_set(x_14, 0, x_32); +return x_14; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; -x_21 = lean_ctor_get(x_11, 0); -x_22 = lean_ctor_get(x_11, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_11); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -x_24 = 0; -lean_inc(x_10); -x_25 = l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrAux(x_23, x_24, x_10); -if (x_25 == 0) +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_14, 1); +lean_inc(x_33); +lean_dec(x_14); +x_34 = l_Lean_Expr_isAnyType(x_15); +lean_dec(x_15); +if (x_34 == 0) { -lean_object* x_26; lean_object* x_27; -lean_dec(x_10); -lean_dec(x_9); +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 = l_Lean_indentExpr(x_1); +x_36 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2; +x_37 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_39 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3(x_39, x_2, x_3, x_4, x_5, x_6, x_33); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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_22); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_box(0); -x_29 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__4(x_10, x_9, x_28, x_2, x_3, x_4, x_5, x_6, x_22); -return x_29; -} -} +return x_40; } else { -lean_object* x_30; lean_object* x_31; -x_30 = l_Lean_Compiler_Simp_findLambda_x3f(x_8, x_4, x_5, x_6, x_7); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -lean_dec(x_9); +lean_object* x_41; lean_object* x_42; lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -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); -lean_dec(x_33); -x_34 = lean_box(0); -lean_ctor_set(x_30, 0, x_34); -return x_30; +lean_dec(x_1); +x_41 = l_Lean_levelOne; +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_33); +return x_42; +} } -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_dec(x_30); -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_35); -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_30, 1); -lean_inc(x_38); -lean_dec(x_30); -x_39 = lean_ctor_get(x_31, 0); -lean_inc(x_39); -lean_dec(x_31); -lean_inc(x_2); -x_40 = l_Lean_Compiler_Simp_shouldInline(x_39, x_2, x_3, x_4, x_5, x_6, x_38); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_unbox(x_41); -lean_dec(x_41); -if (x_42 == 0) -{ uint8_t x_43; -lean_dec(x_39); -lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_43 = !lean_is_exclusive(x_40); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_14); if (x_43 == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_40, 0); -lean_dec(x_44); -x_45 = lean_box(0); -lean_ctor_set(x_40, 0, x_45); -return x_40; +return x_14; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_dec(x_40); -x_47 = lean_box(0); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_14, 0); +x_45 = lean_ctor_get(x_14, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_14); +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_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_40, 1); -lean_inc(x_49); -lean_dec(x_40); -x_50 = lean_box(0); -x_51 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3(x_39, x_9, x_50, x_2, x_3, x_4, x_5, x_6, x_49); -lean_dec(x_39); -return x_51; } } } +static lean_object* _init_l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("lcUnreachable", 13); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2() { _start: { -lean_object* x_10; -x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f___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_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_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_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___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_9; -x_9 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_8; +lean_inc(x_1); +x_8 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2(x_1, x_2, 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; 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_8, 0); +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_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2; +x_14 = l_Lean_Expr_const___override(x_13, x_12); +x_15 = l_Lean_Expr_app___override(x_14, x_1); +lean_ctor_set(x_8, 0, x_15); +return x_8; +} +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_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2; +x_21 = l_Lean_Expr_const___override(x_20, x_19); +x_22 = l_Lean_Expr_app___override(x_21, x_1); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_17); +return x_23; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_3); +uint8_t x_24; lean_dec(x_1); -return x_10; -} +x_24 = !lean_is_exclusive(x_8); +if (x_24 == 0) +{ +return x_8; } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineCandidate_x3f___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: +else { -lean_object* x_10; -x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_3); -return x_10; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_8, 0); +x_26 = lean_ctor_get(x_8, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_8); +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_Lean_Compiler_Simp_expandTrivialExpr___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; -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_Lean_Compiler_Simp_expandTrivialExpr(lean_object* x_1, lean_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_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___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, 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_8; -x_8 = l_Lean_Expr_isFVar(x_1); -if (x_8 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_6, x_5); +if (x_14 == 0) { -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_7); -return x_9; +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_3); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_7); +lean_ctor_set(x_15, 1, x_13); +return x_15; } else { -uint8_t x_10; lean_object* x_11; uint8_t x_12; -x_10 = 1; -lean_inc(x_1); -x_11 = l_Lean_Compiler_Simp_findExpr(x_1, x_10, x_4, x_5, x_6, x_7); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -x_15 = l_Lean_Expr_isLambda(x_13); -if (x_15 == 0) -{ -uint8_t x_16; -x_16 = lean_expr_eqv(x_1, x_13); -lean_dec(x_1); -if (x_16 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; 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_7); +x_16 = lean_array_uget(x_4, x_6); +x_26 = lean_st_ref_get(x_12, x_13); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_get(x_10, x_27); +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); +lean_inc(x_12); +lean_inc(x_11); +x_32 = l_Lean_Compiler_InferType_inferType(x_16, x_31, x_11, x_12, x_30); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_17; uint8_t x_18; -lean_free_object(x_11); -x_17 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_14); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 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_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Compiler_isEmptyType(x_33, x_11, x_12, x_34); +lean_dec(x_33); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) { -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -lean_ctor_set(x_17, 0, x_13); -return x_17; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +lean_inc(x_3); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_3); +x_17 = x_39; +x_18 = x_38; +goto block_25; } 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_13); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} +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_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = lean_expr_instantiate_rev(x_2, x_1); +x_42 = lean_st_ref_get(x_12, x_40); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_st_ref_get(x_10, x_43); +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_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +lean_inc(x_12); +lean_inc(x_11); +x_48 = l_Lean_Compiler_InferType_inferType(x_41, x_47, x_11, x_12, x_46); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +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); +lean_inc(x_12); +lean_inc(x_11); +x_51 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1(x_49, x_8, x_9, x_10, x_11, x_12, x_50); +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; lean_object* x_59; lean_object* x_60; +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 = l_Lean_Compiler_mkLambda(x_1, x_52, x_10, x_11, x_12, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_55); +x_58 = lean_box(0); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_60, 0, x_59); +x_17 = x_60; +x_18 = x_56; +goto block_25; } else { -return x_11; -} +uint8_t x_61; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_3); +x_61 = !lean_is_exclusive(x_51); +if (x_61 == 0) +{ +return x_51; } else { -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_1); -return x_11; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_51, 0); +x_63 = lean_ctor_get(x_51, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_51); +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; +} } } else { -lean_object* x_22; lean_object* x_23; uint8_t 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); +uint8_t x_65; +lean_dec(x_12); lean_dec(x_11); -x_24 = l_Lean_Expr_isLambda(x_22); -if (x_24 == 0) +lean_dec(x_3); +x_65 = !lean_is_exclusive(x_48); +if (x_65 == 0) { -uint8_t x_25; -x_25 = lean_expr_eqv(x_1, x_22); -lean_dec(x_1); -if (x_25 == 0) +return x_48; +} +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = l_Lean_Compiler_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_23); -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); +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_48, 0); +x_67 = lean_ctor_get(x_48, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_48); +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; } -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_28; } -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_27); -return x_29; +} +} +else +{ +uint8_t x_69; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_3); +x_69 = !lean_is_exclusive(x_32); +if (x_69 == 0) +{ +return x_32; } else { -lean_object* x_30; -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_22); -lean_ctor_set(x_30, 1, x_23); -return x_30; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_32, 0); +x_71 = lean_ctor_get(x_32, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_32); +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; +} } +block_25: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_3); +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_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } else { -lean_object* x_31; -lean_dec(x_22); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set(x_31, 1, x_23); -return x_31; +lean_object* x_21; size_t x_22; size_t x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = 1; +x_23 = lean_usize_add(x_6, x_22); +x_6 = x_23; +x_7 = x_21; +x_13 = x_18; +goto _start; } } } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___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_Compiler_Simp_visitLambda___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_9; -x_9 = l_Lean_Compiler_Simp_expandTrivialExpr___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +x_10 = l_Lean_Compiler_Simp_visitLet(x_1, x_2, 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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_13 = l_Lean_Compiler_mkLetUsingScope(x_11, x_6, x_7, x_8, x_12); +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_Compiler_mkLambda(x_2, x_14, x_6, x_7, x_8, 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_2); -return x_9; -} +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_expandTrivialExpr___boxed(lean_object* x_1, 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; -x_8 = l_Lean_Compiler_Simp_expandTrivialExpr(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_17; +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_8; -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___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: +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) { -lean_object* x_9; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_9 = lean_apply_6(x_1, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +return x_13; +} +else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -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_apply_7(x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_11); -return x_12; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +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 = 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 { -uint8_t x_13; +uint8_t 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); -x_13 = !lean_is_exclusive(x_9); -if (x_13 == 0) +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) { -return x_9; +return x_10; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -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_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_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_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Compiler_Simp_visitLambda___closed__1() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1___rarg), 8, 0); +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_Compiler_withNewScope___at_Lean_Compiler_Simp_visitLambda___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 lean_object* l_Lean_Compiler_Simp_visitLambda(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_8; lean_object* x_9; -x_8 = lean_apply_2(x_1, x_2, x_3); -x_9 = l_Lean_Compiler_withNewScopeImp___rarg(x_8, x_4, x_5, x_6, x_7); -return x_9; -} +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_42; lean_object* x_43; 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_9 = lean_st_ref_get(x_7, x_8); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_st_ref_get(x_5, 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_70 = lean_st_ref_get(x_7, x_13); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = lean_st_ref_take(x_5, x_71); +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 = !lean_is_exclusive(x_73); +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; +x_76 = lean_ctor_get(x_73, 1); +lean_dec(x_76); +x_77 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +lean_ctor_set(x_73, 1, x_77); +x_78 = lean_st_ref_set(x_5, x_73, x_74); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = l_Lean_Compiler_visitLambdaCore(x_1, x_5, x_6, x_7, x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (x_2 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_ctor_get(x_81, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_dec(x_81); +x_85 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_86 = l_Lean_Compiler_Simp_visitLambda___lambda__1(x_84, x_83, x_85, x_3, x_4, x_5, x_6, x_7, x_82); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; +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_14 = x_87; +x_15 = x_88; +goto block_41; } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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: +else { -lean_object* x_8; -x_8 = l_Lean_Compiler_visitLambda(x_1, x_4, x_5, x_6, x_7); -return x_8; +lean_object* x_89; lean_object* x_90; +x_89 = lean_ctor_get(x_86, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_86, 1); +lean_inc(x_90); +lean_dec(x_86); +x_42 = x_89; +x_43 = x_90; +goto block_69; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -lean_dec(x_1); -x_10 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; size_t x_95; size_t x_96; lean_object* x_97; lean_object* x_98; +x_91 = lean_ctor_get(x_80, 1); +lean_inc(x_91); +lean_dec(x_80); +x_92 = lean_ctor_get(x_81, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_81, 1); +lean_inc(x_93); +lean_dec(x_81); +x_94 = lean_array_get_size(x_92); +x_95 = lean_usize_of_nat(x_94); +lean_dec(x_94); +x_96 = 0; +x_97 = l_Lean_Compiler_Simp_visitLambda___closed__1; +lean_inc(x_7); lean_inc(x_6); +x_98 = l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___spec__4(x_92, x_93, x_97, x_92, x_95, x_96, x_97, x_3, x_4, x_5, x_6, x_7, x_91); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +lean_dec(x_99); +if (lean_obj_tag(x_100) == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_98, 1); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_box(0); +lean_inc(x_7); lean_inc(x_5); -lean_inc(x_4); -x_11 = l_Lean_Compiler_Simp_visitLet(x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_11) == 0) +x_103 = l_Lean_Compiler_Simp_visitLambda___lambda__1(x_93, x_92, x_102, x_3, x_4, x_5, x_6, x_7, x_101); +if (lean_obj_tag(x_103) == 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); +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +x_14 = x_104; +x_15 = x_105; +goto block_41; +} +else +{ +lean_object* x_106; lean_object* x_107; +x_106 = lean_ctor_get(x_103, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_103, 1); +lean_inc(x_107); +lean_dec(x_103); +x_42 = x_106; +x_43 = x_107; +goto block_69; +} +} +else +{ +lean_object* x_108; lean_object* x_109; +lean_dec(x_93); +lean_dec(x_92); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_108 = lean_ctor_get(x_98, 1); +lean_inc(x_108); +lean_dec(x_98); +x_109 = lean_ctor_get(x_100, 0); +lean_inc(x_109); +lean_dec(x_100); +x_14 = x_109; +x_15 = x_108; +goto block_41; +} +} +else +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_93); +lean_dec(x_92); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_110 = lean_ctor_get(x_98, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_98, 1); +lean_inc(x_111); +lean_dec(x_98); +x_42 = x_110; +x_43 = x_111; +goto block_69; +} +} +} +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; +x_112 = lean_ctor_get(x_73, 0); +x_113 = lean_ctor_get(x_73, 2); +lean_inc(x_113); +lean_inc(x_112); +lean_dec(x_73); +x_114 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_113); +x_116 = lean_st_ref_set(x_5, x_115, x_74); +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +x_118 = l_Lean_Compiler_visitLambdaCore(x_1, x_5, x_6, x_7, x_117); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +if (x_2 == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_ctor_get(x_119, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_dec(x_119); +x_123 = lean_box(0); +lean_inc(x_7); +lean_inc(x_5); +x_124 = l_Lean_Compiler_Simp_visitLambda___lambda__1(x_122, x_121, x_123, x_3, x_4, x_5, x_6, x_7, x_120); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +x_14 = x_125; +x_15 = x_126; +goto block_41; +} +else +{ +lean_object* x_127; lean_object* x_128; +x_127 = lean_ctor_get(x_124, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_124, 1); +lean_inc(x_128); +lean_dec(x_124); +x_42 = x_127; +x_43 = x_128; +goto block_69; +} +} +else +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; size_t x_133; size_t x_134; lean_object* x_135; lean_object* x_136; +x_129 = lean_ctor_get(x_118, 1); +lean_inc(x_129); +lean_dec(x_118); +x_130 = lean_ctor_get(x_119, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_119, 1); +lean_inc(x_131); +lean_dec(x_119); +x_132 = lean_array_get_size(x_130); +x_133 = lean_usize_of_nat(x_132); +lean_dec(x_132); +x_134 = 0; +x_135 = l_Lean_Compiler_Simp_visitLambda___closed__1; +lean_inc(x_7); lean_inc(x_6); +x_136 = l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___spec__4(x_130, x_131, x_135, x_130, x_133, x_134, x_135, x_3, x_4, x_5, x_6, x_7, x_129); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +lean_dec(x_137); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_136, 1); +lean_inc(x_139); +lean_dec(x_136); +x_140 = lean_box(0); +lean_inc(x_7); lean_inc(x_5); -lean_inc(x_4); -x_14 = l_Lean_Compiler_mkLetUsingScope(x_12, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_14) == 0) +x_141 = l_Lean_Compiler_Simp_visitLambda___lambda__1(x_131, x_130, x_140, x_3, x_4, x_5, x_6, x_7, x_139); +if (lean_obj_tag(x_141) == 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_Lean_Compiler_mkLambda(x_8, x_15, x_4, x_5, x_6, x_16); +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +x_14 = x_142; +x_15 = x_143; +goto block_41; +} +else +{ +lean_object* x_144; lean_object* x_145; +x_144 = lean_ctor_get(x_141, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_141, 1); +lean_inc(x_145); +lean_dec(x_141); +x_42 = x_144; +x_43 = x_145; +goto block_69; +} +} +else +{ +lean_object* x_146; lean_object* x_147; +lean_dec(x_131); +lean_dec(x_130); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -lean_dec(x_8); -return x_17; +lean_dec(x_3); +x_146 = lean_ctor_get(x_136, 1); +lean_inc(x_146); +lean_dec(x_136); +x_147 = lean_ctor_get(x_138, 0); +lean_inc(x_147); +lean_dec(x_138); +x_14 = x_147; +x_15 = x_146; +goto block_41; +} } else { -uint8_t x_18; -lean_dec(x_8); +lean_object* x_148; lean_object* x_149; +lean_dec(x_131); +lean_dec(x_130); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) -{ -return x_14; +lean_dec(x_3); +x_148 = lean_ctor_get(x_136, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_136, 1); +lean_inc(x_149); +lean_dec(x_136); +x_42 = x_148; +x_43 = x_149; +goto block_69; } -else +} +} +block_41: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -x_20 = lean_ctor_get(x_14, 1); -lean_inc(x_20); +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_16 = lean_st_ref_get(x_7, x_15); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_get(x_5, x_17); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_14); -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; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_12, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_dec(x_12); +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_24 = lean_ctor_get(x_19, 1); +lean_dec(x_24); +x_25 = lean_ctor_get(x_19, 0); +lean_dec(x_25); +lean_ctor_set(x_19, 1, x_22); +lean_ctor_set(x_19, 0, x_21); +x_26 = lean_st_ref_get(x_7, x_20); +lean_dec(x_7); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_st_ref_set(x_5, x_19, x_27); +lean_dec(x_5); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_14); +return x_28; } +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_14); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } else { -uint8_t x_22; -lean_dec(x_8); -lean_dec(x_6); +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_33 = lean_ctor_get(x_19, 2); +lean_inc(x_33); +lean_dec(x_19); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_21); +lean_ctor_set(x_34, 1, x_22); +lean_ctor_set(x_34, 2, x_33); +x_35 = lean_st_ref_get(x_7, x_20); +lean_dec(x_7); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_st_ref_set(x_5, x_34, x_36); lean_dec(x_5); -lean_dec(x_4); -x_22 = !lean_is_exclusive(x_11); -if (x_22 == 0) +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; +} else { + lean_dec_ref(x_37); + 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_14); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +block_69: { -return x_11; +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_st_ref_get(x_7, x_43); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_st_ref_get(x_5, x_45); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ctor_get(x_12, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_12, 1); +lean_inc(x_50); +lean_dec(x_12); +x_51 = !lean_is_exclusive(x_47); +if (x_51 == 0) +{ +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_47, 1); +lean_dec(x_52); +x_53 = lean_ctor_get(x_47, 0); +lean_dec(x_53); +lean_ctor_set(x_47, 1, x_50); +lean_ctor_set(x_47, 0, x_49); +x_54 = lean_st_ref_get(x_7, x_48); +lean_dec(x_7); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_set(x_5, x_47, x_55); +lean_dec(x_5); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +lean_ctor_set_tag(x_56, 1); +lean_ctor_set(x_56, 0, x_42); +return x_56; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_11, 0); -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_11); -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_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_42); +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; +x_61 = lean_ctor_get(x_47, 2); +lean_inc(x_61); +lean_dec(x_47); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_49); +lean_ctor_set(x_62, 1, x_50); +lean_ctor_set(x_62, 2, x_61); +x_63 = lean_st_ref_get(x_7, x_48); +lean_dec(x_7); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = lean_st_ref_set(x_5, x_62, x_64); +lean_dec(x_5); +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(1, 2, 0); +} else { + x_68 = x_67; + lean_ctor_set_tag(x_68, 1); } -static lean_object* _init_l_Lean_Compiler_Simp_visitLambda___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_visitLambda___lambda__2), 7, 0); -return x_1; +lean_ctor_set(x_68, 0, x_42); +lean_ctor_set(x_68, 1, x_66); +return x_68; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda(lean_object* x_1, 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; -x_8 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_visitLambda___lambda__1___boxed), 7, 1); -lean_closure_set(x_8, 0, x_1); -x_9 = l_Lean_Compiler_Simp_visitLambda___closed__1; -x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1___rarg), 8, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_visitLambda___spec__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7); -return x_11; } } LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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) { @@ -6558,6 +9702,15 @@ return x_38; } } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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 = lean_box(0); +x_12 = lean_apply_8(x_1, x_2, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet(lean_object* x_1, lean_object* x_2, 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: { @@ -6579,35 +9732,159 @@ lean_dec(x_11); x_15 = l_Lean_Expr_isLambda(x_14); if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l_Lean_Compiler_Simp_visitLet___lambda__1(x_12, x_2, x_10, x_9, x_13, x_14, x_16, x_3, x_4, x_5, x_6, x_7, x_8); -return x_17; +lean_object* x_16; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_14); +x_16 = l_Lean_Compiler_Simp_simpValue_x3f(x_14, x_3, x_4, x_5, x_6, x_7, x_8); +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) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +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_Compiler_Simp_visitLet___lambda__1(x_12, x_2, x_10, x_9, x_13, x_14, x_19, x_3, x_4, x_5, x_6, x_7, x_18); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_14); +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = lean_ctor_get(x_17, 0); +lean_inc(x_22); +lean_dec(x_17); +x_23 = l_Lean_Expr_isLet(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Lean_Compiler_Simp_visitLet___lambda__1(x_12, x_2, x_10, x_9, x_13, x_22, x_24, x_3, x_4, x_5, x_6, x_7, x_21); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_unsigned_to_nat(0u); +x_27 = l_Lean_Compiler_Simp_mkFlatLet_go(x_9, x_10, x_12, x_13, x_22, x_26); +lean_dec(x_12); +x_28 = l_Lean_Compiler_Simp_visitLet(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +return x_28; +} +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(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +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; +} +} +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +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_37 = !lean_is_exclusive(x_16); +if (x_37 == 0) +{ +return x_16; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_16, 0); +x_39 = lean_ctor_get(x_16, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_16); +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_18; +uint8_t x_41; lean_object* x_42; +x_41 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_18 = l_Lean_Compiler_Simp_visitLambda(x_14, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_18) == 0) +x_42 = l_Lean_Compiler_Simp_visitLambda(x_14, x_41, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* 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_box(0); -x_22 = l_Lean_Compiler_Simp_visitLet___lambda__1(x_12, x_2, x_10, x_9, x_13, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_20); -return x_22; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +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_box(0); +x_46 = l_Lean_Compiler_Simp_visitLet___lambda__1(x_12, x_2, x_10, x_9, x_13, x_43, x_45, x_3, x_4, x_5, x_6, x_7, x_44); +return x_46; } else { -uint8_t x_23; +uint8_t x_47; lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); @@ -6617,181 +9894,245 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +x_47 = !lean_is_exclusive(x_42); +if (x_47 == 0) { -return x_18; +return x_42; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -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_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_42, 0); +x_49 = lean_ctor_get(x_42, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_42); +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_27; lean_object* x_28; -x_27 = lean_expr_instantiate_rev(x_1, x_2); +lean_object* x_51; lean_object* x_52; +x_51 = lean_expr_instantiate_rev(x_1, x_2); lean_dec(x_2); lean_dec(x_1); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_27); -x_28 = l_Lean_Compiler_isCasesApp_x3f(x_27, x_6, x_7, x_8); -if (lean_obj_tag(x_28) == 0) +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_51); +x_52 = l_Lean_Compiler_Simp_simpValue_x3f(x_51, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_29; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -if (lean_obj_tag(x_29) == 0) +lean_object* x_53; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_box(0); -x_32 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_51); +x_55 = l_Lean_Compiler_isCasesApp_x3f(x_51, x_6, x_7, x_54); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +if (lean_obj_tag(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_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_box(0); +x_59 = l_Lean_Compiler_Simp_internalize_visitLambda___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_27); -x_33 = l_Lean_Compiler_Simp_inlineApp_x3f(x_27, x_32, x_31, x_3, x_4, x_5, x_6, x_7, x_30); -if (lean_obj_tag(x_33) == 0) +lean_inc(x_51); +x_60 = l_Lean_Compiler_Simp_inlineApp_x3f(x_51, x_59, x_58, x_3, x_4, x_5, x_6, x_7, x_57); +if (lean_obj_tag(x_60) == 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_61; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_obj_tag(x_61) == 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 = l_Lean_Compiler_Simp_expandTrivialExpr(x_27, x_3, x_4, x_5, x_6, x_7, x_35); +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_Compiler_Simp_expandTrivialExpr(x_51, x_3, x_4, x_5, x_6, x_7, x_62); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_36; +return x_63; } else { -uint8_t x_37; -lean_dec(x_27); +uint8_t x_64; +lean_dec(x_51); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_37 = !lean_is_exclusive(x_33); -if (x_37 == 0) +x_64 = !lean_is_exclusive(x_60); +if (x_64 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_33, 0); -lean_dec(x_38); -x_39 = lean_ctor_get(x_34, 0); -lean_inc(x_39); -lean_dec(x_34); -lean_ctor_set(x_33, 0, x_39); -return x_33; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_60, 0); +lean_dec(x_65); +x_66 = lean_ctor_get(x_61, 0); +lean_inc(x_66); +lean_dec(x_61); +lean_ctor_set(x_60, 0, x_66); +return x_60; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -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); -lean_dec(x_34); -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_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_60, 1); +lean_inc(x_67); +lean_dec(x_60); +x_68 = lean_ctor_get(x_61, 0); +lean_inc(x_68); +lean_dec(x_61); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_67); +return x_69; } } } else { -uint8_t x_43; -lean_dec(x_27); +uint8_t x_70; +lean_dec(x_51); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_43 = !lean_is_exclusive(x_33); -if (x_43 == 0) +x_70 = !lean_is_exclusive(x_60); +if (x_70 == 0) { -return x_33; +return x_60; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_33, 0); -x_45 = lean_ctor_get(x_33, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_33); -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_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_60, 0); +x_72 = lean_ctor_get(x_60, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_60); +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_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 1); -lean_inc(x_47); -lean_dec(x_28); -x_48 = lean_ctor_get(x_29, 0); -lean_inc(x_48); -lean_dec(x_29); -x_49 = l_Lean_Compiler_Simp_visitCases(x_48, x_27, x_3, x_4, x_5, x_6, x_7, x_47); -return x_49; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_55, 1); +lean_inc(x_74); +lean_dec(x_55); +x_75 = lean_ctor_get(x_56, 0); +lean_inc(x_75); +lean_dec(x_56); +x_76 = l_Lean_Compiler_Simp_visitCases(x_75, x_51, x_3, x_4, x_5, x_6, x_7, x_74); +return x_76; } } else { -uint8_t x_50; -lean_dec(x_27); +uint8_t x_77; +lean_dec(x_51); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_50 = !lean_is_exclusive(x_28); -if (x_50 == 0) +x_77 = !lean_is_exclusive(x_55); +if (x_77 == 0) { -return x_28; +return x_55; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_28, 0); -x_52 = lean_ctor_get(x_28, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_28); -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; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_55, 0); +x_79 = lean_ctor_get(x_55, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_55); +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_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_51); +x_81 = lean_ctor_get(x_52, 1); +lean_inc(x_81); +lean_dec(x_52); +x_82 = lean_ctor_get(x_53, 0); +lean_inc(x_82); +lean_dec(x_53); +x_83 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_1 = x_82; +x_2 = x_83; +x_8 = x_81; +goto _start; +} +} +else +{ +uint8_t x_85; +lean_dec(x_51); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_85 = !lean_is_exclusive(x_52); +if (x_85 == 0) +{ +return x_52; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_52, 0); +x_87 = lean_ctor_get(x_52, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_52); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } } @@ -6827,36 +10168,37 @@ goto _start; } 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; lean_object* x_23; uint8_t x_24; lean_object* x_25; x_21 = lean_array_fget(x_5, x_2); x_22 = lean_box(0); x_23 = lean_array_fset(x_5, x_2, x_22); +x_24 = 1; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_24 = l_Lean_Compiler_Simp_visitLambda(x_21, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_24) == 0) +x_25 = l_Lean_Compiler_Simp_visitLambda(x_21, x_24, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_array_fset(x_23, x_2, x_25); -x_28 = lean_nat_add(x_2, x_4); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_array_fset(x_23, x_2, x_26); +x_29 = lean_nat_add(x_2, x_4); lean_dec(x_2); x_1 = x_16; -x_2 = x_28; -x_5 = x_27; -x_11 = x_26; +x_2 = x_29; +x_5 = x_28; +x_11 = x_27; goto _start; } else { -uint8_t x_30; +uint8_t x_31; lean_dec(x_23); lean_dec(x_16); lean_dec(x_10); @@ -6865,30 +10207,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_24); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_25); +if (x_31 == 0) { -return x_24; +return x_25; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_24, 0); -x_32 = lean_ctor_get(x_24, 1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_25, 0); +x_33 = lean_ctor_get(x_25, 1); +lean_inc(x_33); lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_24); -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_dec(x_25); +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; } } } } else { -lean_object* x_34; +lean_object* x_35; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6896,15 +10238,15 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_5); -lean_ctor_set(x_34, 1, x_11); -return x_34; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_5); +lean_ctor_set(x_35, 1, x_11); +return x_35; } } else { -lean_object* x_35; +lean_object* x_36; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6912,46 +10254,28 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_5); -lean_ctor_set(x_35, 1, x_11); -return x_35; -} -} -} -static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Core_instMonadCoreM; -x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); -return x_2; -} +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_5); +lean_ctor_set(x_36, 1, x_11); +return x_36; } -static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1; -x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); -return x_2; } } -static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3() { +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2; +x_1 = l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2; x_2 = l_Lean_instInhabitedExpr; x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4() { +static lean_object* _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3; +x_1 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1; x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -6961,7 +10285,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(le _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4; +x_8 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2; x_9 = lean_panic_fn(x_8, x_1); x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; @@ -6997,25 +10321,17 @@ static lean_object* _init_l_Lean_Compiler_Simp_visitCases___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp", 18); -return x_1; -} -} -static lean_object* _init_l_Lean_Compiler_Simp_visitCases___closed__5() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp.visitCases", 29); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_visitCases___closed__6() { +static lean_object* _init_l_Lean_Compiler_Simp_visitCases___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_Lean_Compiler_Simp_visitCases___closed__4; -x_2 = l_Lean_Compiler_Simp_visitCases___closed__5; -x_3 = lean_unsigned_to_nat(179u); +x_1 = l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1; +x_2 = l_Lean_Compiler_Simp_visitCases___closed__4; +x_3 = lean_unsigned_to_nat(414u); x_4 = lean_unsigned_to_nat(4u); x_5 = l_Lean_Compiler_Simp_visitCases___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7025,643 +10341,713 @@ return x_6; LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitCases(lean_object* x_1, lean_object* x_2, 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; lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_9); -x_11 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_10); -x_12 = lean_mk_array(x_10, x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_10, x_13); -lean_dec(x_10); -lean_inc(x_2); -x_15 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_12, x_14); -x_16 = lean_ctor_get(x_1, 2); -lean_inc(x_16); -x_17 = lean_ctor_get(x_16, 1); +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; lean_object* x_22; +x_9 = l_Lean_Expr_getAppFn(x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_10); +x_12 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_11); +x_13 = lean_mk_array(x_11, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_11, x_14); +lean_dec(x_11); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_13, x_15); +x_17 = lean_ctor_get(x_1, 2); lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_nat_sub(x_17, x_13); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); lean_dec(x_17); -x_19 = lean_array_get_size(x_15); -x_20 = lean_nat_dec_lt(x_18, x_19); -if (x_20 == 0) -{ -lean_object* x_81; lean_object* x_82; +x_19 = lean_nat_sub(x_18, x_14); lean_dec(x_18); -x_81 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_82 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_81); -x_21 = x_82; -goto block_80; +x_20 = lean_array_get_size(x_16); +x_21 = lean_nat_dec_lt(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_90; lean_object* x_91; +lean_dec(x_19); +x_90 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_91 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_90); +x_22 = x_91; +goto block_89; } else { -lean_object* x_83; -x_83 = lean_array_fget(x_15, x_18); -lean_dec(x_18); -x_21 = x_83; -goto block_80; +lean_object* x_92; +x_92 = lean_array_fget(x_16, x_19); +lean_dec(x_19); +x_22 = x_92; +goto block_89; } -block_80: +block_89: { -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; -x_22 = 1; -x_23 = l_Lean_Compiler_Simp_findExpr(x_21, x_22, x_5, x_6, x_7, x_8); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +uint8_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; lean_object* x_31; +x_23 = 1; +x_24 = l_Lean_Compiler_Simp_findExpr(x_22, x_23, x_5, x_6, x_7, x_8); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_st_ref_get(x_7, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_st_ref_get(x_7, x_26); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = l_Lean_Expr_constructorApp_x3f(x_29, x_24); -if (lean_obj_tag(x_30) == 0) +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Expr_constructorApp_x3f(x_30, x_25); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_19); -x_31 = lean_ctor_get(x_1, 3); -lean_inc(x_31); -lean_dec(x_1); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 0); +lean_object* x_32; +lean_dec(x_20); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_16); +lean_inc(x_9); +lean_inc(x_1); +x_32 = l_Lean_Compiler_Simp_simpCasesOnCases_x3f(x_1, x_9, x_16, x_3, x_4, x_5, x_6, x_7, x_29); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 2); +if (lean_obj_tag(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; +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -lean_dec(x_31); -lean_inc(x_32); -x_35 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_visitCases___spec__1(x_32, x_33, x_32, x_34, x_15, x_3, x_4, x_5, x_6, x_7, x_28); -lean_dec(x_34); lean_dec(x_32); -if (lean_obj_tag(x_35) == 0) +x_35 = lean_ctor_get(x_1, 3); +lean_inc(x_35); +lean_dec(x_1); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_35, 2); +lean_inc(x_38); +lean_dec(x_35); +lean_inc(x_36); +x_39 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_visitCases___spec__1(x_36, x_37, x_36, x_38, x_16, x_3, x_4, x_5, x_6, x_7, x_34); +lean_dec(x_38); +lean_dec(x_36); +if (lean_obj_tag(x_39) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_35, 0); -x_38 = l_Lean_Expr_getAppFn(x_2); -lean_dec(x_2); -x_39 = l_Lean_mkAppN(x_38, x_37); -lean_ctor_set(x_35, 0, x_39); -return x_35; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_39, 0); +x_42 = l_Lean_mkAppN(x_9, x_41); +lean_ctor_set(x_39, 0, x_42); +return x_39; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_40 = lean_ctor_get(x_35, 0); -x_41 = lean_ctor_get(x_35, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_35); -x_42 = l_Lean_Expr_getAppFn(x_2); -lean_dec(x_2); -x_43 = l_Lean_mkAppN(x_42, x_40); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); -return x_44; +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 = l_Lean_mkAppN(x_9, 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_44); +return x_46; } } else { -uint8_t x_45; -lean_dec(x_2); -x_45 = !lean_is_exclusive(x_35); -if (x_45 == 0) +uint8_t x_47; +lean_dec(x_9); +x_47 = !lean_is_exclusive(x_39); +if (x_47 == 0) { -return x_35; +return x_39; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_35, 0); -x_47 = lean_ctor_get(x_35, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_35); -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_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; } } } 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_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -lean_dec(x_2); -x_49 = lean_ctor_get(x_30, 0); -lean_inc(x_49); -lean_dec(x_30); -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_dec(x_16); +lean_dec(x_9); +x_51 = lean_ctor_get(x_32, 1); lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 2); +lean_dec(x_32); +x_52 = lean_ctor_get(x_33, 0); lean_inc(x_52); -x_53 = lean_ctor_get(x_1, 3); -lean_inc(x_53); -lean_dec(x_1); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -lean_dec(x_53); -x_55 = lean_nat_add(x_54, x_52); -lean_dec(x_52); -lean_dec(x_54); -x_56 = lean_nat_dec_lt(x_55, x_19); -lean_dec(x_19); -x_57 = lean_ctor_get(x_50, 3); -lean_inc(x_57); -lean_dec(x_50); -x_58 = lean_array_get_size(x_51); -x_59 = l_Array_toSubarray___rarg(x_51, x_57, x_58); -x_60 = l_Array_ofSubarray___rarg(x_59); -if (x_56 == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -lean_dec(x_55); -lean_dec(x_15); -x_61 = l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4; -x_62 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_61); -x_63 = l_Lean_Expr_beta(x_62, x_60); -x_64 = l_Lean_Expr_isLambda(x_63); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = l_Lean_Compiler_Simp_markSimplified___rarg(x_4, x_5, x_6, x_7, x_28); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_68 = l_Lean_Compiler_Simp_visitLet(x_63, x_67, x_3, x_4, x_5, x_6, x_7, x_66); -return x_68; -} -else -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_63); -x_69 = l_Lean_Compiler_Simp_visitCases___closed__6; -x_70 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_69, x_3, x_4, x_5, x_6, x_7, x_28); -return x_70; +lean_dec(x_33); +x_2 = x_52; +x_8 = x_51; +goto _start; } } else { -lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_71 = lean_array_fget(x_15, x_55); -lean_dec(x_55); -lean_dec(x_15); -x_72 = l_Lean_Expr_beta(x_71, x_60); -x_73 = l_Lean_Expr_isLambda(x_72); -if (x_73 == 0) +uint8_t x_54; +lean_dec(x_16); +lean_dec(x_9); +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_54 = !lean_is_exclusive(x_32); +if (x_54 == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_74 = l_Lean_Compiler_Simp_markSimplified___rarg(x_4, x_5, x_6, x_7, x_28); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -lean_dec(x_74); -x_76 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_77 = l_Lean_Compiler_Simp_visitLet(x_72, x_76, x_3, x_4, x_5, x_6, x_7, x_75); -return x_77; +return x_32; } else { -lean_object* x_78; lean_object* x_79; -lean_dec(x_72); -x_78 = l_Lean_Compiler_Simp_visitCases___closed__6; -x_79 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_78, x_3, x_4, x_5, x_6, x_7, x_28); -return x_79; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = 0; -x_10 = l_Lean_Compiler_mkLocalDecl(x_1, x_2, x_9, x_5, x_6, x_7, x_8); -return x_10; -} -} -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_32, 0); +x_56 = lean_ctor_get(x_32, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_32); +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; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_9 = l_Lean_Compiler_mkLetUsingScope(x_2, 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; -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_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1; -x_13 = lean_array_push(x_12, x_1); -x_14 = l_Lean_Compiler_mkLambda(x_13, x_10, x_5, x_6, x_7, x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_13); -return x_14; } else { -uint8_t x_15; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +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; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_9); +x_58 = lean_ctor_get(x_31, 0); +lean_inc(x_58); +lean_dec(x_31); +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, 2); +lean_inc(x_61); +x_62 = lean_ctor_get(x_1, 3); +lean_inc(x_62); lean_dec(x_1); -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +lean_dec(x_62); +x_64 = lean_nat_add(x_63, x_61); +lean_dec(x_61); +lean_dec(x_63); +x_65 = lean_nat_dec_lt(x_64, x_20); +lean_dec(x_20); +x_66 = lean_ctor_get(x_59, 3); +lean_inc(x_66); +lean_dec(x_59); +x_67 = lean_array_get_size(x_60); +x_68 = l_Array_toSubarray___rarg(x_60, x_66, x_67); +x_69 = l_Array_ofSubarray___rarg(x_68); +if (x_65 == 0) { -return x_9; +lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +lean_dec(x_64); +lean_dec(x_16); +x_70 = l_Lean_Compiler_Simp_simpProj_x3f___closed__4; +x_71 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_70); +x_72 = l_Lean_Expr_beta(x_71, x_69); +x_73 = l_Lean_Expr_isLambda(x_72); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = l_Lean_Compiler_Simp_markSimplified___rarg(x_4, x_5, x_6, x_7, x_29); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_77 = l_Lean_Compiler_Simp_visitLet(x_72, x_76, x_3, x_4, x_5, x_6, x_7, x_75); +return x_77; } else { -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_9, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_9); -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_object* x_78; lean_object* x_79; +lean_dec(x_72); +x_78 = l_Lean_Compiler_Simp_visitCases___closed__5; +x_79 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_78, x_3, x_4, x_5, x_6, x_7, x_29); +return x_79; } } +else +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = lean_array_fget(x_16, x_64); +lean_dec(x_64); +lean_dec(x_16); +x_81 = l_Lean_Expr_beta(x_80, x_69); +x_82 = l_Lean_Expr_isLambda(x_81); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = l_Lean_Compiler_Simp_markSimplified___rarg(x_4, x_5, x_6, x_7, x_29); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_86 = l_Lean_Compiler_Simp_visitLet(x_81, x_85, x_3, x_4, x_5, x_6, x_7, x_84); +return x_86; } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_x", 2); -return x_1; +lean_object* x_87; lean_object* x_88; +lean_dec(x_81); +x_87 = l_Lean_Compiler_Simp_visitCases___closed__5; +x_88 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_87, x_3, x_4, x_5, x_6, x_7, x_29); +return x_88; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_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_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); -return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4() { +static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Option.get!", 11); +x_1 = l_Lean_inheritedTraceOptions; return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5() { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("value is none", 13); -return x_1; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1; +x_11 = lean_st_ref_get(x_10, 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; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_5, 2); +x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_13, x_14, x_1); +lean_dec(x_13); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6() { -_start: +else { -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_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3; -x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_5, 2); +x_20 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_17, x_19, x_1); +lean_dec(x_17); +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_18); +return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_1, x_13); -lean_dec(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; -x_15 = lean_array_get_size(x_2); -x_16 = l_Array_toSubarray___rarg(x_2, x_3, x_15); -x_17 = l_Array_ofSubarray___rarg(x_16); -lean_inc(x_6); -x_18 = l_Lean_mkAppN(x_6, x_17); -x_19 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; -lean_inc(x_11); -lean_inc(x_10); -x_20 = l_Lean_Compiler_mkAuxLetDecl(x_18, x_19, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_20) == 0) -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* 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_inc(x_21); -x_23 = lean_array_push(x_5, x_21); +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; uint8_t x_27; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_st_ref_get(x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_24 = l_Lean_Compiler_Simp_visitLet(x_21, x_23, x_7, x_8, x_9, x_10, x_11, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_get(x_7, x_12); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_5, x_15); +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 = lean_ctor_get(x_6, 2); +x_21 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6; +lean_inc(x_20); +x_22 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_19); +lean_ctor_set(x_22, 3, x_20); +x_23 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_2); +x_24 = lean_st_ref_take(x_7, x_18); 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_Compiler_Simp_inlineApp_x3f___lambda__2(x_6, x_25, x_7, x_8, x_9, x_10, x_11, x_26); -lean_dec(x_8); -lean_dec(x_7); -return x_27; +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_28 = lean_ctor_get(x_25, 3); +x_29 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_30 = 0; +x_31 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_23); +lean_ctor_set(x_31, 2, x_29); +lean_ctor_set_uint8(x_31, sizeof(void*)*3, x_30); +lean_inc(x_9); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_9); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Std_PersistentArray_push___rarg(x_28, x_32); +lean_ctor_set(x_25, 3, x_33); +x_34 = lean_st_ref_set(x_7, x_25, x_26); +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 = lean_box(0); +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 = lean_box(0); +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; 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; 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_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +x_43 = lean_ctor_get(x_25, 2); +x_44 = lean_ctor_get(x_25, 3); +x_45 = lean_ctor_get(x_25, 4); +x_46 = lean_ctor_get(x_25, 5); +x_47 = lean_ctor_get(x_25, 6); +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_25); +x_48 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_49 = 0; +x_50 = lean_alloc_ctor(12, 3, 1); +lean_ctor_set(x_50, 0, x_1); +lean_ctor_set(x_50, 1, x_23); +lean_ctor_set(x_50, 2, x_48); +lean_ctor_set_uint8(x_50, sizeof(void*)*3, x_49); +lean_inc(x_9); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_9); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Std_PersistentArray_push___rarg(x_44, x_51); +x_53 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_53, 0, x_41); +lean_ctor_set(x_53, 1, x_42); +lean_ctor_set(x_53, 2, x_43); +lean_ctor_set(x_53, 3, x_52); +lean_ctor_set(x_53, 4, x_45); +lean_ctor_set(x_53, 5, x_46); +lean_ctor_set(x_53, 6, x_47); +x_54 = lean_st_ref_set(x_7, x_53, x_26); +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_54); + x_56 = lean_box(0); +} +x_57 = lean_box(0); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_56; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; } -else -{ -uint8_t 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); -x_28 = !lean_is_exclusive(x_24); -if (x_28 == 0) -{ -return x_24; } -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_24, 0); -x_30 = lean_ctor_get(x_24, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_24); -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; } +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_32 = lean_ctor_get(x_20, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_20, 1); -lean_inc(x_33); -lean_dec(x_20); -x_34 = lean_ctor_get(x_4, 0); -lean_inc(x_34); -lean_dec(x_4); -x_35 = lean_array_push(x_5, x_32); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_9; lean_inc(x_7); -x_36 = l_Lean_Compiler_Simp_visitLet(x_34, x_35, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_6); +lean_inc(x_5); +x_9 = l_Lean_Compiler_mkLetUsingScope(x_2, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -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); -x_39 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_6, x_37, x_7, x_8, x_9, x_10, x_11, x_38); -lean_dec(x_8); +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_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_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1; +x_13 = lean_array_push(x_12, x_1); +x_14 = l_Lean_Compiler_mkLambda(x_13, x_10, x_5, x_6, x_7, x_11); lean_dec(x_7); -return x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_13); +return x_14; } else { -uint8_t x_40; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +uint8_t x_15; lean_dec(x_7); lean_dec(x_6); -x_40 = !lean_is_exclusive(x_36); -if (x_40 == 0) +lean_dec(x_5); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) { -return x_36; +return x_9; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_36, 0); -x_42 = lean_ctor_get(x_36, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_36); -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_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_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +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_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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: { -uint8_t x_44; -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); -x_44 = !lean_is_exclusive(x_20); -if (x_44 == 0) +lean_object* x_10; +x_10 = l_Lean_Compiler_Simp_visitLet(x_2, x_1, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -return x_20; +uint8_t x_11; +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); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_10, 0, x_13); +return x_10; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_20, 0); -x_46 = lean_ctor_get(x_20, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_20); -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_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); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_14); +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_48; -lean_dec(x_3); -lean_dec(x_2); -lean_inc(x_6); -x_48 = lean_array_push(x_5, x_6); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6; -x_50 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_49); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_51 = l_Lean_Compiler_Simp_visitLet(x_50, x_48, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_51) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_10); +if (x_18 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -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 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_6, x_52, x_7, x_8, x_9, x_10, x_11, x_53); -lean_dec(x_8); -lean_dec(x_7); -return x_54; +return x_10; } else { -uint8_t x_55; -lean_dec(x_11); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_10, 0); +x_20 = lean_ctor_get(x_10, 1); +lean_inc(x_20); +lean_inc(x_19); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 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); +return x_21; +} +} +} +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1() { +_start: { -return x_51; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_x", 2); +return x_1; } -else +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2() { +_start: { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_51); -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_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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_dec(x_5); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_3); +x_12 = lean_box(0); +x_13 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_1, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } else { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_4, 0); -lean_inc(x_59); -lean_dec(x_4); -lean_inc(x_11); +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 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_st_ref_get(x_10, x_11); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(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); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_60 = l_Lean_Compiler_Simp_visitLet(x_59, x_48, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_60) == 0) +x_21 = l_Lean_Compiler_InferType_inferType(x_3, x_20, x_9, x_10, x_19); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_6, x_61, x_7, x_8, x_9, x_10, x_11, x_62); -lean_dec(x_8); -lean_dec(x_7); -return x_63; +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; lean_object* x_32; +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 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; +x_25 = l_Lean_Compiler_mkAuxLetDeclName(x_24, x_8, x_9, x_10, x_23); +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 = 0; +x_29 = lean_unsigned_to_nat(0u); +x_30 = l_Lean_Compiler_Simp_mkFlatLet_go(x_26, x_22, x_14, x_28, x_4, x_29); +lean_dec(x_14); +x_31 = lean_box(0); +x_32 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_1, x_30, x_31, x_6, x_7, x_8, x_9, x_10, x_27); +return x_32; } else -{ -uint8_t x_64; -lean_dec(x_11); +{ +uint8_t x_33; +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_64 = !lean_is_exclusive(x_60); -if (x_64 == 0) +lean_dec(x_4); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) { -return x_60; +return x_21; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_60, 0); -x_66 = lean_ctor_get(x_60, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_60); -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_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_21, 0); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_21); +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; } } } } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__1() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -7669,1046 +11055,1571 @@ x_1 = lean_mk_string_from_bytes("_y", 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__2() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__1; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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, lean_object* x_9) { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3() { _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_1); -x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f(x_1, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4() { +_start: { -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Option.get!", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5() { +_start: { -uint8_t x_12; -lean_dec(x_8); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("value is none", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___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_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_x3f___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) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_7); +x_14 = l_Lean_Compiler_Simp_markSimplified___rarg(x_9, x_10, x_11, x_12, x_13); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_340; uint8_t x_341; +x_340 = lean_ctor_get(x_1, 0); +lean_inc(x_340); +x_341 = lean_nat_dec_eq(x_4, x_340); +lean_dec(x_340); +if (x_341 == 0) +{ +lean_object* x_342; +x_342 = lean_box(0); +x_16 = x_342; +goto block_339; +} +else +{ +lean_object* x_343; lean_object* x_344; lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); +x_343 = lean_ctor_get(x_1, 1); +lean_inc(x_343); lean_dec(x_1); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_344 = l_Lean_Compiler_Simp_betaReduce(x_343, x_3, x_8, x_9, x_10, x_11, x_12, x_15); +if (lean_obj_tag(x_344) == 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_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); +x_346 = lean_ctor_get(x_344, 1); +lean_inc(x_346); +lean_dec(x_344); +x_347 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_348 = l_Lean_Compiler_Simp_visitLet(x_345, x_347, x_8, x_9, x_10, x_11, x_12, x_346); +if (lean_obj_tag(x_348) == 0) { -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; -} +uint8_t x_349; +x_349 = !lean_is_exclusive(x_348); +if (x_349 == 0) +{ +lean_object* x_350; lean_object* x_351; +x_350 = lean_ctor_get(x_348, 0); +x_351 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_351, 0, x_350); +lean_ctor_set(x_348, 0, x_351); +return x_348; } 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; -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_ctor_get(x_11, 0); -lean_inc(x_19); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - x_20 = x_11; -} else { - lean_dec_ref(x_11); - x_20 = lean_box(0); +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +x_352 = lean_ctor_get(x_348, 0); +x_353 = lean_ctor_get(x_348, 1); +lean_inc(x_353); +lean_inc(x_352); +lean_dec(x_348); +x_354 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_354, 0, x_352); +x_355 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_355, 0, x_354); +lean_ctor_set(x_355, 1, x_353); +return x_355; } -x_21 = lean_unsigned_to_nat(0u); -x_22 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_21); -x_23 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; -lean_inc(x_22); -x_24 = lean_mk_array(x_22, x_23); -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_sub(x_22, x_25); -lean_dec(x_22); -lean_inc(x_1); -x_27 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_24, x_26); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_83; -x_83 = lean_nat_dec_eq(x_19, x_21); -if (x_83 == 0) -{ -lean_object* x_84; -x_84 = lean_box(0); -x_28 = x_84; -goto block_82; } else { -lean_object* x_85; lean_object* x_86; -lean_dec(x_27); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_2); -x_85 = lean_box(0); -x_86 = l_Lean_Compiler_Simp_inlineApp(x_1, x_85, x_4, x_5, x_6, x_7, x_8, x_18); -if (lean_obj_tag(x_86) == 0) -{ -uint8_t x_87; -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) +uint8_t x_356; +x_356 = !lean_is_exclusive(x_348); +if (x_356 == 0) { -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_86, 0); -x_89 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_86, 0, x_89); -return x_86; +return x_348; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_86, 0); -x_91 = lean_ctor_get(x_86, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_86); -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_90); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_91); -return x_93; +lean_object* x_357; lean_object* x_358; lean_object* x_359; +x_357 = lean_ctor_get(x_348, 0); +x_358 = lean_ctor_get(x_348, 1); +lean_inc(x_358); +lean_inc(x_357); +lean_dec(x_348); +x_359 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_359, 0, x_357); +lean_ctor_set(x_359, 1, x_358); +return x_359; +} } } else { -uint8_t x_94; -x_94 = !lean_is_exclusive(x_86); -if (x_94 == 0) +uint8_t x_360; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_360 = !lean_is_exclusive(x_344); +if (x_360 == 0) { -return x_86; +return x_344; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_86, 0); -x_96 = lean_ctor_get(x_86, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_86); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_object* x_361; lean_object* x_362; lean_object* x_363; +x_361 = lean_ctor_get(x_344, 0); +x_362 = lean_ctor_get(x_344, 1); +lean_inc(x_362); +lean_inc(x_361); +lean_dec(x_344); +x_363 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_363, 0, x_361); +lean_ctor_set(x_363, 1, x_362); +return x_363; } } } } else { -lean_object* x_98; -x_98 = lean_box(0); -x_28 = x_98; -goto block_82; +lean_object* x_364; +x_364 = lean_box(0); +x_16 = x_364; +goto block_339; } -block_82: +block_339: { -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_dec(x_28); -x_29 = l_Lean_Expr_getAppFn(x_1); +lean_object* x_17; lean_object* x_18; +lean_dec(x_16); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_17); +x_18 = l_Lean_Compiler_onlyOneExitPoint(x_17, x_11, x_12, x_15); +if (lean_obj_tag(x_18) == 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_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_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = l_Lean_Expr_getAppFn(x_2); +lean_dec(x_2); +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); lean_dec(x_1); -x_30 = lean_array_get_size(x_27); -x_31 = lean_nat_sub(x_30, x_19); -lean_dec(x_30); +x_24 = lean_unsigned_to_nat(0u); +lean_inc(x_23); +lean_inc(x_3); +x_25 = l_Array_toSubarray___rarg(x_3, x_24, x_23); +x_26 = l_Array_ofSubarray___rarg(x_25); +lean_inc(x_26); +x_27 = l_Lean_mkAppN(x_22, x_26); +x_28 = lean_st_ref_get(x_12, x_21); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_st_ref_get(x_10, x_29); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_inc(x_27); -x_32 = l_Array_toSubarray___rarg(x_27, x_21, x_31); -x_33 = l_Array_ofSubarray___rarg(x_32); -x_34 = l_Lean_mkAppN(x_29, x_33); -x_35 = lean_st_ref_get(x_8, x_18); -x_36 = lean_ctor_get(x_35, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +lean_inc(x_12); +lean_inc(x_11); +x_34 = l_Lean_Compiler_InferType_inferType(x_27, x_33, x_11, x_12, x_32); +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; lean_object* x_41; 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_116; lean_object* x_117; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; +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_35); -x_37 = lean_st_ref_get(x_6, x_36); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); +lean_dec(x_34); +x_37 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2; +x_38 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_37, x_11, x_12, x_36); +x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); lean_dec(x_38); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_34); -x_41 = l_Lean_Compiler_InferType_inferType(x_34, x_40, x_7, x_8, x_39); -if (lean_obj_tag(x_41) == 0) +x_83 = lean_st_ref_get(x_12, x_40); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = lean_st_ref_get(x_10, x_84); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_144 = lean_st_ref_get(x_12, x_87); +x_145 = lean_ctor_get(x_144, 1); +lean_inc(x_145); +lean_dec(x_144); +x_146 = lean_st_ref_take(x_10, x_145); +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +x_149 = !lean_is_exclusive(x_147); +if (x_149 == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_150 = lean_ctor_get(x_147, 1); +lean_dec(x_150); +x_151 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +lean_ctor_set(x_147, 1, x_151); +x_152 = lean_st_ref_set(x_10, x_147, x_148); +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +lean_dec(x_152); +x_154 = 0; +x_155 = l_Lean_Compiler_mkLocalDecl(x_39, x_35, x_154, x_10, x_11, x_12, 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); +lean_dec(x_155); +x_158 = lean_nat_dec_eq(x_4, x_23); +lean_dec(x_4); +if (x_158 == 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_51; -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 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__2; -x_45 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_44, x_7, x_8, x_43); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___boxed), 8, 2); -lean_closure_set(x_48, 0, x_46); -lean_closure_set(x_48, 1, x_42); -x_49 = lean_alloc_closure((void*)(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3), 12, 5); -lean_closure_set(x_49, 0, x_19); -lean_closure_set(x_49, 1, x_27); -lean_closure_set(x_49, 2, x_31); -lean_closure_set(x_49, 3, x_3); -lean_closure_set(x_49, 4, x_2); -x_50 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_Simp_visitLambda___spec__1___rarg), 8, 2); -lean_closure_set(x_50, 0, x_48); -lean_closure_set(x_50, 1, x_49); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_51 = l_Lean_Compiler_withNewScope___at_Lean_Compiler_Simp_visitLambda___spec__2(x_50, x_4, x_5, x_6, x_7, x_8, x_47); -if (lean_obj_tag(x_51) == 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; +x_159 = lean_array_get_size(x_3); +x_160 = l_Array_toSubarray___rarg(x_3, x_23, x_159); +x_161 = l_Array_ofSubarray___rarg(x_160); +lean_inc(x_156); +x_162 = l_Lean_mkAppN(x_156, x_161); +x_163 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; +lean_inc(x_12); +lean_inc(x_11); +x_164 = l_Lean_Compiler_mkAuxLetDecl(x_162, x_163, x_10, x_11, x_12, x_157); +if (lean_obj_tag(x_164) == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -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); -lean_inc(x_8); -lean_inc(x_7); -x_54 = l_Lean_Compiler_mkJpDeclIfNotSimple(x_52, x_6, x_7, x_8, x_53); -if (lean_obj_tag(x_54) == 0) +if (lean_obj_tag(x_5) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -if (lean_is_scalar(x_20)) { - x_57 = lean_alloc_ctor(1, 1, 0); -} else { - x_57 = x_20; -} -lean_ctor_set(x_57, 0, x_55); -x_58 = l_Lean_Compiler_Simp_inlineApp(x_34, x_57, x_4, x_5, x_6, x_7, x_8, x_56); -if (lean_obj_tag(x_58) == 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); +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +lean_inc(x_165); +x_167 = lean_array_push(x_6, x_165); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_168 = l_Lean_Compiler_Simp_visitLet(x_165, x_167, x_8, x_9, x_10, x_11, x_12, x_166); +if (lean_obj_tag(x_168) == 0) { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_168, 1); +lean_inc(x_170); +lean_dec(x_168); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_171 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_156, x_169, x_8, x_9, x_10, x_11, x_12, x_170); +if (lean_obj_tag(x_171) == 0) { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_58, 0); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_58, 0, x_61); -return x_58; +lean_object* x_172; lean_object* x_173; +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_88 = x_172; +x_89 = x_173; +goto block_115; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_62 = lean_ctor_get(x_58, 0); -x_63 = lean_ctor_get(x_58, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_58); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_62); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -return x_65; -} +lean_object* x_174; lean_object* x_175; +x_174 = lean_ctor_get(x_171, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_171, 1); +lean_inc(x_175); +lean_dec(x_171); +x_116 = x_174; +x_117 = x_175; +goto block_143; } -else -{ -uint8_t x_66; -x_66 = !lean_is_exclusive(x_58); -if (x_66 == 0) -{ -return x_58; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_58, 0); -x_68 = lean_ctor_get(x_58, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_58); -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_object* x_176; lean_object* x_177; +lean_dec(x_156); +x_176 = lean_ctor_get(x_168, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_168, 1); +lean_inc(x_177); +lean_dec(x_168); +x_116 = x_176; +x_117 = x_177; +goto block_143; } } else { -uint8_t x_70; -lean_dec(x_34); -lean_dec(x_20); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_164, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_164, 1); +lean_inc(x_179); +lean_dec(x_164); +x_180 = lean_ctor_get(x_5, 0); +lean_inc(x_180); lean_dec(x_5); -lean_dec(x_4); -x_70 = !lean_is_exclusive(x_54); -if (x_70 == 0) +x_181 = lean_array_push(x_6, x_178); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_182 = l_Lean_Compiler_Simp_visitLet(x_180, x_181, x_8, x_9, x_10, x_11, x_12, x_179); +if (lean_obj_tag(x_182) == 0) { -return x_54; -} -else +lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_182, 1); +lean_inc(x_184); +lean_dec(x_182); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_185 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_156, x_183, x_8, x_9, x_10, x_11, x_12, x_184); +if (lean_obj_tag(x_185) == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_54, 0); -x_72 = lean_ctor_get(x_54, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_54); -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; -} -} +lean_object* x_186; lean_object* x_187; +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_88 = x_186; +x_89 = x_187; +goto block_115; } else { -uint8_t x_74; -lean_dec(x_34); -lean_dec(x_20); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_74 = !lean_is_exclusive(x_51); -if (x_74 == 0) -{ -return x_51; +lean_object* x_188; lean_object* x_189; +x_188 = lean_ctor_get(x_185, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_185, 1); +lean_inc(x_189); +lean_dec(x_185); +x_116 = x_188; +x_117 = x_189; +goto block_143; +} } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_51, 0); -x_76 = lean_ctor_get(x_51, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_51); -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; +lean_object* x_190; lean_object* x_191; +lean_dec(x_156); +x_190 = lean_ctor_get(x_182, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_182, 1); +lean_inc(x_191); +lean_dec(x_182); +x_116 = x_190; +x_117 = x_191; +goto block_143; } } } else { -uint8_t x_78; -lean_dec(x_34); -lean_dec(x_31); -lean_dec(x_27); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_192; lean_object* x_193; +lean_dec(x_156); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +x_192 = lean_ctor_get(x_164, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_164, 1); +lean_inc(x_193); +lean_dec(x_164); +x_116 = x_192; +x_117 = x_193; +goto block_143; +} +} +else +{ +lean_object* x_194; +lean_dec(x_23); lean_dec(x_3); -lean_dec(x_2); -x_78 = !lean_is_exclusive(x_41); -if (x_78 == 0) +lean_inc(x_156); +x_194 = lean_array_push(x_6, x_156); +if (lean_obj_tag(x_5) == 0) { -return x_41; +lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_195 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6; +x_196 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_195); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_197 = l_Lean_Compiler_Simp_visitLet(x_196, x_194, x_8, x_9, x_10, x_11, x_12, x_157); +if (lean_obj_tag(x_197) == 0) +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_198 = lean_ctor_get(x_197, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_197, 1); +lean_inc(x_199); +lean_dec(x_197); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_200 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_156, x_198, x_8, x_9, x_10, x_11, x_12, x_199); +if (lean_obj_tag(x_200) == 0) +{ +lean_object* x_201; lean_object* x_202; +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +lean_dec(x_200); +x_88 = x_201; +x_89 = x_202; +goto block_115; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_41, 0); -x_80 = lean_ctor_get(x_41, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_41); -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_object* x_203; lean_object* x_204; +x_203 = lean_ctor_get(x_200, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_200, 1); +lean_inc(x_204); +lean_dec(x_200); +x_116 = x_203; +x_117 = x_204; +goto block_143; } } +else +{ +lean_object* x_205; lean_object* x_206; +lean_dec(x_156); +x_205 = lean_ctor_get(x_197, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_197, 1); +lean_inc(x_206); +lean_dec(x_197); +x_116 = x_205; +x_117 = x_206; +goto block_143; } } else { -uint8_t x_99; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_207; lean_object* x_208; +x_207 = lean_ctor_get(x_5, 0); +lean_inc(x_207); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_99 = !lean_is_exclusive(x_10); -if (x_99 == 0) +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_208 = l_Lean_Compiler_Simp_visitLet(x_207, x_194, x_8, x_9, x_10, x_11, x_12, x_157); +if (lean_obj_tag(x_208) == 0) { -return x_10; +lean_object* x_209; lean_object* x_210; lean_object* x_211; +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); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_211 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_156, x_209, x_8, x_9, x_10, x_11, x_12, x_210); +if (lean_obj_tag(x_211) == 0) +{ +lean_object* x_212; lean_object* x_213; +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_88 = x_212; +x_89 = x_213; +goto block_115; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_10, 0); -x_101 = lean_ctor_get(x_10, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_10); -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; -} -} -} +lean_object* x_214; lean_object* x_215; +x_214 = lean_ctor_get(x_211, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_211, 1); +lean_inc(x_215); +lean_dec(x_211); +x_116 = x_214; +x_117 = x_215; +goto block_143; } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___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; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = lean_st_ref_get(x_6, x_7); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -lean_dec(x_8); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1; -x_11 = lean_st_ref_get(x_10, 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; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_5, 2); -x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_13, x_14, x_1); -lean_dec(x_13); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_ctor_get(x_5, 2); -x_20 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_17, x_19, x_1); -lean_dec(x_17); -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_18); -return x_22; +lean_object* x_216; lean_object* x_217; +lean_dec(x_156); +x_216 = lean_ctor_get(x_208, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_208, 1); +lean_inc(x_217); +lean_dec(x_208); +x_116 = x_216; +x_117 = x_217; +goto block_143; } } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp___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: +} +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; 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; uint8_t x_27; -x_9 = lean_ctor_get(x_6, 5); -x_10 = lean_st_ref_get(x_7, 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); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_st_ref_get(x_7, x_12); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_st_ref_get(x_5, x_15); -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 = lean_ctor_get(x_6, 2); -x_21 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6; -lean_inc(x_20); -x_22 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_22, 0, x_13); -lean_ctor_set(x_22, 1, x_21); -lean_ctor_set(x_22, 2, x_19); -lean_ctor_set(x_22, 3, x_20); -x_23 = lean_alloc_ctor(6, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_2); -x_24 = lean_st_ref_take(x_7, x_18); -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_is_exclusive(x_25); -if (x_27 == 0) +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; uint8_t x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; uint8_t x_228; +x_218 = lean_ctor_get(x_147, 0); +x_219 = lean_ctor_get(x_147, 2); +lean_inc(x_219); +lean_inc(x_218); +lean_dec(x_147); +x_220 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_221 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_221, 0, x_218); +lean_ctor_set(x_221, 1, x_220); +lean_ctor_set(x_221, 2, x_219); +x_222 = lean_st_ref_set(x_10, x_221, x_148); +x_223 = lean_ctor_get(x_222, 1); +lean_inc(x_223); +lean_dec(x_222); +x_224 = 0; +x_225 = l_Lean_Compiler_mkLocalDecl(x_39, x_35, x_224, x_10, x_11, x_12, x_223); +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_225, 1); +lean_inc(x_227); +lean_dec(x_225); +x_228 = lean_nat_dec_eq(x_4, x_23); +lean_dec(x_4); +if (x_228 == 0) { -lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_28 = lean_ctor_get(x_25, 3); -x_29 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_30 = 0; -x_31 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set(x_31, 1, x_23); -lean_ctor_set(x_31, 2, x_29); -lean_ctor_set_uint8(x_31, sizeof(void*)*3, x_30); +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_229 = lean_array_get_size(x_3); +x_230 = l_Array_toSubarray___rarg(x_3, x_23, x_229); +x_231 = l_Array_ofSubarray___rarg(x_230); +lean_inc(x_226); +x_232 = l_Lean_mkAppN(x_226, x_231); +x_233 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; +lean_inc(x_12); +lean_inc(x_11); +x_234 = l_Lean_Compiler_mkAuxLetDecl(x_232, x_233, x_10, x_11, x_12, x_227); +if (lean_obj_tag(x_234) == 0) +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); +lean_inc(x_236); +lean_dec(x_234); +lean_inc(x_235); +x_237 = lean_array_push(x_6, x_235); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_9); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Std_PersistentArray_push___rarg(x_28, x_32); -lean_ctor_set(x_25, 3, x_33); -x_34 = lean_st_ref_set(x_7, x_25, x_26); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_inc(x_8); +x_238 = l_Lean_Compiler_Simp_visitLet(x_235, x_237, x_8, x_9, x_10, x_11, x_12, x_236); +if (lean_obj_tag(x_238) == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -x_37 = lean_box(0); -lean_ctor_set(x_34, 0, x_37); -return x_34; +lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_238, 1); +lean_inc(x_240); +lean_dec(x_238); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_241 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_226, x_239, x_8, x_9, x_10, x_11, x_12, x_240); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; lean_object* x_243; +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_88 = x_242; +x_89 = x_243; +goto block_115; } 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 = lean_box(0); -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; +lean_object* x_244; lean_object* x_245; +x_244 = lean_ctor_get(x_241, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_241, 1); +lean_inc(x_245); +lean_dec(x_241); +x_116 = x_244; +x_117 = x_245; +goto block_143; } } 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; 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_41 = lean_ctor_get(x_25, 0); -x_42 = lean_ctor_get(x_25, 1); -x_43 = lean_ctor_get(x_25, 2); -x_44 = lean_ctor_get(x_25, 3); -x_45 = lean_ctor_get(x_25, 4); -x_46 = lean_ctor_get(x_25, 5); -x_47 = lean_ctor_get(x_25, 6); -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_25); -x_48 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_49 = 0; -x_50 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_50, 0, x_1); -lean_ctor_set(x_50, 1, x_23); -lean_ctor_set(x_50, 2, x_48); -lean_ctor_set_uint8(x_50, sizeof(void*)*3, x_49); +lean_object* x_246; lean_object* x_247; +lean_dec(x_226); +x_246 = lean_ctor_get(x_238, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_238, 1); +lean_inc(x_247); +lean_dec(x_238); +x_116 = x_246; +x_117 = x_247; +goto block_143; +} +} +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_248 = lean_ctor_get(x_234, 0); +lean_inc(x_248); +x_249 = lean_ctor_get(x_234, 1); +lean_inc(x_249); +lean_dec(x_234); +x_250 = lean_ctor_get(x_5, 0); +lean_inc(x_250); +lean_dec(x_5); +x_251 = lean_array_push(x_6, x_248); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_9); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Std_PersistentArray_push___rarg(x_44, x_51); -x_53 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_53, 0, x_41); -lean_ctor_set(x_53, 1, x_42); -lean_ctor_set(x_53, 2, x_43); -lean_ctor_set(x_53, 3, x_52); -lean_ctor_set(x_53, 4, x_45); -lean_ctor_set(x_53, 5, x_46); -lean_ctor_set(x_53, 6, x_47); -x_54 = lean_st_ref_set(x_7, x_53, x_26); -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_54); - x_56 = lean_box(0); +lean_inc(x_8); +x_252 = l_Lean_Compiler_Simp_visitLet(x_250, x_251, x_8, x_9, x_10, x_11, x_12, x_249); +if (lean_obj_tag(x_252) == 0) +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; +x_253 = lean_ctor_get(x_252, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_252, 1); +lean_inc(x_254); +lean_dec(x_252); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_255 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_226, x_253, x_8, x_9, x_10, x_11, x_12, x_254); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_255, 1); +lean_inc(x_257); +lean_dec(x_255); +x_88 = x_256; +x_89 = x_257; +goto block_115; } -x_57 = lean_box(0); -if (lean_is_scalar(x_56)) { - x_58 = lean_alloc_ctor(0, 2, 0); -} else { - x_58 = x_56; +else +{ +lean_object* x_258; lean_object* x_259; +x_258 = lean_ctor_get(x_255, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_255, 1); +lean_inc(x_259); +lean_dec(x_255); +x_116 = x_258; +x_117 = x_259; +goto block_143; } -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_55); -return x_58; +} +else +{ +lean_object* x_260; lean_object* x_261; +lean_dec(x_226); +x_260 = lean_ctor_get(x_252, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_252, 1); +lean_inc(x_261); +lean_dec(x_252); +x_116 = x_260; +x_117 = x_261; +goto block_143; } } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("!value.isLambda\n ", 18); -return x_1; +lean_object* x_262; lean_object* x_263; +lean_dec(x_226); +lean_dec(x_6); +lean_dec(x_5); +x_262 = lean_ctor_get(x_234, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_234, 1); +lean_inc(x_263); +lean_dec(x_234); +x_116 = x_262; +x_117 = x_263; +goto block_143; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_visitCases___closed__1; -x_2 = l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_264; +lean_dec(x_23); +lean_dec(x_3); +lean_inc(x_226); +x_264 = lean_array_push(x_6, x_226); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6; +x_266 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_265); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_267 = l_Lean_Compiler_Simp_visitLet(x_266, x_264, x_8, x_9, x_10, x_11, x_12, x_227); +if (lean_obj_tag(x_267) == 0) +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_267, 1); +lean_inc(x_269); +lean_dec(x_267); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_270 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_226, x_268, x_8, x_9, x_10, x_11, x_12, x_269); +if (lean_obj_tag(x_270) == 0) +{ +lean_object* x_271; lean_object* x_272; +x_271 = lean_ctor_get(x_270, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_270, 1); +lean_inc(x_272); +lean_dec(x_270); +x_88 = x_271; +x_89 = x_272; +goto block_115; +} +else +{ +lean_object* x_273; lean_object* x_274; +x_273 = lean_ctor_get(x_270, 0); +lean_inc(x_273); +x_274 = lean_ctor_get(x_270, 1); +lean_inc(x_274); +lean_dec(x_270); +x_116 = x_273; +x_117 = x_274; +goto block_143; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Compiler.Simp.inlineApp", 28); -return x_1; +lean_object* x_275; lean_object* x_276; +lean_dec(x_226); +x_275 = lean_ctor_get(x_267, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_267, 1); +lean_inc(x_276); +lean_dec(x_267); +x_116 = x_275; +x_117 = x_276; +goto block_143; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4() { -_start: +else { -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_Compiler_Simp_visitCases___closed__4; -x_2 = l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3; -x_3 = lean_unsigned_to_nat(200u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Compiler_Simp_inlineApp___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_object* x_277; lean_object* x_278; +x_277 = lean_ctor_get(x_5, 0); +lean_inc(x_277); +lean_dec(x_5); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_278 = l_Lean_Compiler_Simp_visitLet(x_277, x_264, x_8, x_9, x_10, x_11, x_12, x_227); +if (lean_obj_tag(x_278) == 0) +{ +lean_object* x_279; lean_object* x_280; lean_object* x_281; +x_279 = lean_ctor_get(x_278, 0); +lean_inc(x_279); +x_280 = lean_ctor_get(x_278, 1); +lean_inc(x_280); +lean_dec(x_278); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_281 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_226, x_279, x_8, x_9, x_10, x_11, x_12, x_280); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_88 = x_282; +x_89 = x_283; +goto block_115; +} +else +{ +lean_object* x_284; lean_object* x_285; +x_284 = lean_ctor_get(x_281, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_281, 1); +lean_inc(x_285); +lean_dec(x_281); +x_116 = x_284; +x_117 = x_285; +goto block_143; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp___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: +else +{ +lean_object* x_286; lean_object* x_287; +lean_dec(x_226); +x_286 = lean_ctor_get(x_278, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_278, 1); +lean_inc(x_287); +lean_dec(x_278); +x_116 = x_286; +x_117 = x_287; +goto block_143; +} +} +} +} +block_82: { -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_unsigned_to_nat(0u); -x_12 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_11); -x_13 = l_Lean_Compiler_Simp_collectInlineStats_go___closed__1; +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +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); lean_inc(x_12); -x_14 = lean_mk_array(x_12, x_13); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_12, x_15); -lean_dec(x_12); -x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_14, x_16); -x_18 = l_Lean_Expr_beta(x_4, x_17); +lean_inc(x_11); +x_44 = l_Lean_Compiler_mkJpDeclIfNotSimple(x_42, x_10, x_11, x_12, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +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); +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 = l_Lean_Compiler_attachOptJp(x_18, x_2, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_19) == 0) +x_47 = l_Lean_Compiler_Simp_betaReduce(x_17, x_26, x_8, x_9, x_10, x_11, x_12, x_46); +if (lean_obj_tag(x_47) == 0) { -lean_object* x_20; lean_object* x_21; uint8_t 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 = l_Lean_Expr_isLambda(x_20); -if (x_22 == 0) +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_50 = l_Lean_Compiler_attachJp(x_48, x_45, x_10, x_11, x_12, x_49); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = l_Lean_Compiler_Simp_markSimplified___rarg(x_6, x_7, x_8, x_9, x_21); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = !lean_is_exclusive(x_5); -if (x_25 == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_54 = l_Lean_Compiler_Simp_visitLet(x_51, x_53, x_8, x_9, x_10, x_11, x_12, x_52); +if (lean_obj_tag(x_54) == 0) +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_54, 0, x_57); +return x_54; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_54, 0); +x_59 = lean_ctor_get(x_54, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_54); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_58); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} +} +else { -uint8_t x_26; -x_26 = l_Lean_Expr_isConst(x_3); -lean_dec(x_3); -if (x_26 == 0) +uint8_t x_62; +x_62 = !lean_is_exclusive(x_54); +if (x_62 == 0) { -uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_27 = 1; -lean_ctor_set_uint8(x_5, sizeof(void*)*2, x_27); -x_28 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_29 = l_Lean_Compiler_Simp_visitLet(x_20, x_28, x_5, x_6, x_7, x_8, x_9, x_24); -return x_29; +return x_54; } else { -uint8_t x_30; lean_object* x_31; lean_object* x_32; -x_30 = 0; -lean_ctor_set_uint8(x_5, sizeof(void*)*2, x_30); -x_31 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_32 = l_Lean_Compiler_Simp_visitLet(x_20, x_31, x_5, x_6, x_7, x_8, x_9, x_24); -return x_32; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_54, 0); +x_64 = lean_ctor_get(x_54, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_54); +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 { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -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_Lean_Expr_isConst(x_3); -lean_dec(x_3); -if (x_35 == 0) +uint8_t x_66; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_66 = !lean_is_exclusive(x_50); +if (x_66 == 0) { -uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = 1; -x_37 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_37, 0, x_33); -lean_ctor_set(x_37, 1, x_34); -lean_ctor_set_uint8(x_37, sizeof(void*)*2, x_36); -x_38 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_39 = l_Lean_Compiler_Simp_visitLet(x_20, x_38, x_37, x_6, x_7, x_8, x_9, x_24); -return x_39; +return x_50; } else { -uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = 0; -x_41 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_41, 0, x_33); -lean_ctor_set(x_41, 1, x_34); -lean_ctor_set_uint8(x_41, sizeof(void*)*2, x_40); -x_42 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_43 = l_Lean_Compiler_Simp_visitLet(x_20, x_42, x_41, x_6, x_7, x_8, x_9, x_24); -return x_43; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_50, 0); +x_68 = lean_ctor_get(x_50, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_50); +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; } } } else { -lean_object* x_44; lean_object* x_45; -lean_dec(x_20); -lean_dec(x_3); -x_44 = l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4; -x_45 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_44, x_5, x_6, x_7, x_8, x_9, x_21); -return x_45; +uint8_t x_70; +lean_dec(x_45); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +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_46; +uint8_t x_74; +lean_dec(x_26); +lean_dec(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_6); -lean_dec(x_5); -lean_dec(x_3); -x_46 = !lean_is_exclusive(x_19); -if (x_46 == 0) +x_74 = !lean_is_exclusive(x_44); +if (x_74 == 0) { -return x_19; +return x_44; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_19, 0); -x_48 = lean_ctor_get(x_19, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_19); -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_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_44, 0); +x_76 = lean_ctor_get(x_44, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_44); +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_26); +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_78 = !lean_is_exclusive(x_41); +if (x_78 == 0) +{ +return x_41; } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); -return x_1; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_41, 0); +x_80 = lean_ctor_get(x_41, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_41); +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; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2() { -_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_Compiler_Simp_visitCases___closed__4; -x_2 = l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3; -x_3 = lean_unsigned_to_nat(195u); -x_4 = lean_unsigned_to_nat(43u); -x_5 = l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; } +block_115: +{ +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; uint8_t x_97; +x_90 = lean_st_ref_get(x_12, x_89); +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +x_92 = lean_st_ref_get(x_10, x_91); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_86, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_86, 1); +lean_inc(x_96); +lean_dec(x_86); +x_97 = !lean_is_exclusive(x_93); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_98 = lean_ctor_get(x_93, 1); +lean_dec(x_98); +x_99 = lean_ctor_get(x_93, 0); +lean_dec(x_99); +lean_ctor_set(x_93, 1, x_96); +lean_ctor_set(x_93, 0, x_95); +x_100 = lean_st_ref_get(x_12, x_94); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +lean_dec(x_100); +x_102 = lean_st_ref_set(x_10, x_93, x_101); +x_103 = !lean_is_exclusive(x_102); +if (x_103 == 0) +{ +lean_object* x_104; +x_104 = lean_ctor_get(x_102, 0); +lean_dec(x_104); +lean_ctor_set(x_102, 0, x_88); +x_41 = x_102; +goto block_82; } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3() { -_start: +else { -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_Compiler_Simp_visitCases___closed__4; -x_2 = l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3; -x_3 = lean_unsigned_to_nat(192u); -x_4 = lean_unsigned_to_nat(48u); -x_5 = l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_102, 1); +lean_inc(x_105); +lean_dec(x_102); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_88); +lean_ctor_set(x_106, 1, x_105); +x_41 = x_106; +goto block_82; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp___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: +else { -lean_dec(x_4); -if (lean_obj_tag(x_3) == 4) +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_107 = lean_ctor_get(x_93, 2); +lean_inc(x_107); +lean_dec(x_93); +x_108 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_108, 0, x_95); +lean_ctor_set(x_108, 1, x_96); +lean_ctor_set(x_108, 2, x_107); +x_109 = lean_st_ref_get(x_12, x_94); +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +lean_dec(x_109); +x_111 = lean_st_ref_set(x_10, x_108, x_110); +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_113 = x_111; +} else { + lean_dec_ref(x_111); + x_113 = lean_box(0); +} +if (lean_is_scalar(x_113)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_113; +} +lean_ctor_set(x_114, 0, x_88); +lean_ctor_set(x_114, 1, x_112); +x_41 = x_114; +goto block_82; +} +} +block_143: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); -lean_inc(x_9); -lean_inc(x_8); -x_13 = l_Lean_Compiler_getStage1Decl_x3f(x_11, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 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; uint8_t x_125; +x_118 = lean_st_ref_get(x_12, x_117); +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +lean_dec(x_118); +x_120 = lean_st_ref_get(x_10, x_119); +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec(x_120); +x_123 = lean_ctor_get(x_86, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_86, 1); +lean_inc(x_124); +lean_dec(x_86); +x_125 = !lean_is_exclusive(x_121); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_126 = lean_ctor_get(x_121, 1); +lean_dec(x_126); +x_127 = lean_ctor_get(x_121, 0); +lean_dec(x_127); +lean_ctor_set(x_121, 1, x_124); +lean_ctor_set(x_121, 0, x_123); +x_128 = lean_st_ref_get(x_12, x_122); +x_129 = lean_ctor_get(x_128, 1); +lean_inc(x_129); +lean_dec(x_128); +x_130 = lean_st_ref_set(x_10, x_121, x_129); +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 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_132; +x_132 = lean_ctor_get(x_130, 0); +lean_dec(x_132); +lean_ctor_set_tag(x_130, 1); +lean_ctor_set(x_130, 0, x_116); +x_41 = x_130; +goto block_82; +} +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_17 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_16, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_17) == 0) +lean_object* x_133; lean_object* x_134; +x_133 = lean_ctor_get(x_130, 1); +lean_inc(x_133); +lean_dec(x_130); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_116); +lean_ctor_set(x_134, 1, x_133); +x_41 = x_134; +goto block_82; +} +} +else { -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_Lean_Compiler_Simp_inlineApp___lambda__1(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_19); -return x_20; +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_135 = lean_ctor_get(x_121, 2); +lean_inc(x_135); +lean_dec(x_121); +x_136 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_136, 0, x_123); +lean_ctor_set(x_136, 1, x_124); +lean_ctor_set(x_136, 2, x_135); +x_137 = lean_st_ref_get(x_12, x_122); +x_138 = lean_ctor_get(x_137, 1); +lean_inc(x_138); +lean_dec(x_137); +x_139 = lean_st_ref_set(x_10, x_136, x_138); +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_139); + x_141 = lean_box(0); +} +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); +} else { + x_142 = x_141; + lean_ctor_set_tag(x_142, 1); +} +lean_ctor_set(x_142, 0, x_116); +lean_ctor_set(x_142, 1, x_140); +x_41 = x_142; +goto block_82; +} +} } else { -uint8_t x_21; +uint8_t x_288; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(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_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_288 = !lean_is_exclusive(x_34); +if (x_288 == 0) { -return x_17; +return x_34; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -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 = 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_289; lean_object* x_290; lean_object* x_291; +x_289 = lean_ctor_get(x_34, 0); +x_290 = lean_ctor_get(x_34, 1); +lean_inc(x_290); +lean_inc(x_289); +lean_dec(x_34); +x_291 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set(x_291, 1, x_290); +return x_291; } } } 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; -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_dec(x_13); -x_26 = lean_ctor_get(x_14, 0); -lean_inc(x_26); -lean_dec(x_14); -x_27 = lean_ctor_get(x_26, 3); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParams___spec__1(x_28, x_12, x_27); -lean_dec(x_12); -lean_dec(x_28); -x_30 = l_Lean_Compiler_Simp_inlineApp___lambda__1(x_1, x_2, x_3, x_29, x_5, x_6, x_7, x_8, x_9, x_25); -return x_30; +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_292 = lean_ctor_get(x_18, 1); +lean_inc(x_292); +lean_dec(x_18); +x_293 = lean_ctor_get(x_1, 0); +lean_inc(x_293); +lean_dec(x_1); +x_294 = lean_unsigned_to_nat(0u); +lean_inc(x_293); +lean_inc(x_3); +x_295 = l_Array_toSubarray___rarg(x_3, x_294, x_293); +x_296 = l_Array_ofSubarray___rarg(x_295); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_296); +x_297 = l_Lean_Compiler_Simp_betaReduce(x_17, x_296, x_8, x_9, x_10, x_11, x_12, x_292); +if (lean_obj_tag(x_297) == 0) +{ +lean_object* x_298; lean_object* x_299; uint8_t x_300; +x_298 = lean_ctor_get(x_297, 0); +lean_inc(x_298); +x_299 = lean_ctor_get(x_297, 1); +lean_inc(x_299); +lean_dec(x_297); +x_300 = lean_nat_dec_lt(x_293, x_4); +lean_dec(x_4); +if (x_300 == 0) +{ +lean_object* x_301; lean_object* x_302; +lean_dec(x_296); +lean_dec(x_293); +lean_dec(x_3); +x_301 = lean_box(0); +x_302 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3(x_6, x_5, x_2, x_298, x_301, x_8, x_9, x_10, x_11, x_12, x_299); +return x_302; } +else +{ +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; +x_303 = l_Lean_Expr_getAppFn(x_2); +x_304 = l_Lean_mkAppN(x_303, x_296); +x_305 = lean_st_ref_get(x_12, x_299); +x_306 = lean_ctor_get(x_305, 1); +lean_inc(x_306); +lean_dec(x_305); +x_307 = lean_st_ref_get(x_10, x_306); +x_308 = lean_ctor_get(x_307, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_307, 1); +lean_inc(x_309); +lean_dec(x_307); +x_310 = lean_ctor_get(x_308, 0); +lean_inc(x_310); +lean_dec(x_308); +lean_inc(x_12); +lean_inc(x_11); +x_311 = l_Lean_Compiler_InferType_inferType(x_304, x_310, x_11, x_12, x_309); +if (lean_obj_tag(x_311) == 0) +{ +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; uint8_t x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; +x_312 = lean_ctor_get(x_311, 0); +lean_inc(x_312); +x_313 = lean_ctor_get(x_311, 1); +lean_inc(x_313); +lean_dec(x_311); +x_314 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2; +x_315 = l_Lean_Compiler_mkAuxLetDeclName(x_314, x_10, x_11, x_12, x_313); +x_316 = lean_ctor_get(x_315, 0); +lean_inc(x_316); +x_317 = lean_ctor_get(x_315, 1); +lean_inc(x_317); +lean_dec(x_315); +x_318 = lean_array_get_size(x_3); +x_319 = l_Array_toSubarray___rarg(x_3, x_293, x_318); +x_320 = l_Array_ofSubarray___rarg(x_319); +x_321 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7; +x_322 = l_Lean_mkAppN(x_321, x_320); +x_323 = 0; +x_324 = l_Lean_Compiler_Simp_mkFlatLet_go(x_316, x_312, x_322, x_323, x_298, x_294); +lean_dec(x_322); +x_325 = lean_box(0); +x_326 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3(x_6, x_5, x_2, x_324, x_325, x_8, x_9, x_10, x_11, x_12, x_317); +return x_326; } else { -uint8_t x_31; +uint8_t x_327; +lean_dec(x_298); +lean_dec(x_293); 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_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_13); -if (x_31 == 0) +x_327 = !lean_is_exclusive(x_311); +if (x_327 == 0) { -return x_13; +return x_311; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_13, 0); -x_33 = lean_ctor_get(x_13, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_13); -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_object* x_328; lean_object* x_329; lean_object* x_330; +x_328 = lean_ctor_get(x_311, 0); +x_329 = lean_ctor_get(x_311, 1); +lean_inc(x_329); +lean_inc(x_328); +lean_dec(x_311); +x_330 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_330, 0, x_328); +lean_ctor_set(x_330, 1, x_329); +return x_330; +} } } } else { -lean_object* x_35; lean_object* x_36; -lean_inc(x_3); -x_35 = l_Lean_Compiler_Simp_findLambda_x3f(x_3, x_7, x_8, x_9, x_10); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) +uint8_t x_331; +lean_dec(x_296); +lean_dec(x_293); +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_4); +lean_dec(x_3); +lean_dec(x_2); +x_331 = !lean_is_exclusive(x_297); +if (x_331 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_39 = l_panic___at_Lean_Compiler_Simp_visitCases___spec__2(x_38, x_5, x_6, x_7, x_8, x_9, x_37); -if (lean_obj_tag(x_39) == 0) +return x_297; +} +else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -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 = l_Lean_Compiler_Simp_inlineApp___lambda__1(x_1, x_2, x_3, x_40, x_5, x_6, x_7, x_8, x_9, x_41); -return x_42; +lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_332 = lean_ctor_get(x_297, 0); +x_333 = lean_ctor_get(x_297, 1); +lean_inc(x_333); +lean_inc(x_332); +lean_dec(x_297); +x_334 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_334, 0, x_332); +lean_ctor_set(x_334, 1, x_333); +return x_334; +} +} +} } else { -uint8_t x_43; +uint8_t x_335; +lean_dec(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_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_43 = !lean_is_exclusive(x_39); -if (x_43 == 0) +x_335 = !lean_is_exclusive(x_18); +if (x_335 == 0) { -return x_39; +return x_18; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_39, 0); -x_45 = lean_ctor_get(x_39, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_39); -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_object* x_336; lean_object* x_337; lean_object* x_338; +x_336 = lean_ctor_get(x_18, 0); +x_337 = lean_ctor_get(x_18, 1); +lean_inc(x_337); +lean_inc(x_336); +lean_dec(x_18); +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; } -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_47 = lean_ctor_get(x_35, 1); -lean_inc(x_47); -lean_dec(x_35); -x_48 = lean_ctor_get(x_36, 0); -lean_inc(x_48); -lean_dec(x_36); -x_49 = l_Lean_LocalDecl_value(x_48); -lean_dec(x_48); -x_50 = l_Lean_Compiler_Simp_inlineApp___lambda__1(x_1, x_2, x_3, x_49, x_5, x_6, x_7, x_8, x_9, x_47); -return x_50; } } } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__1() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__1() { _start: { lean_object* x_1; @@ -8716,17 +12627,17 @@ x_1 = lean_mk_string_from_bytes("Compiler", 8); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__2() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_Simp_inlineApp___closed__1; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__3() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__3() { _start: { lean_object* x_1; @@ -8734,17 +12645,17 @@ x_1 = lean_mk_string_from_bytes("simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__4() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_inlineApp___closed__2; -x_2 = l_Lean_Compiler_Simp_inlineApp___closed__3; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__2; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__5() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__5() { _start: { lean_object* x_1; @@ -8752,17 +12663,17 @@ x_1 = lean_mk_string_from_bytes("inline", 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__6() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_inlineApp___closed__4; -x_2 = l_Lean_Compiler_Simp_inlineApp___closed__5; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__7() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__7() { _start: { lean_object* x_1; @@ -8770,75 +12681,231 @@ x_1 = lean_mk_string_from_bytes("inlining ", 9); return x_1; } } -static lean_object* _init_l_Lean_Compiler_Simp_inlineApp___closed__8() { +static lean_object* _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Compiler_Simp_inlineApp___closed__7; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp(lean_object* x_1, lean_object* x_2, lean_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_Compiler_Simp_inlineApp_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, lean_object* x_9) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_Expr_getAppFn(x_1); -x_10 = l_Lean_Compiler_Simp_inlineApp___closed__6; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___spec__1(x_10, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) +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_1); +x_10 = l_Lean_Compiler_Simp_inlineCandidate_x3f(x_1, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_box(0); -x_16 = l_Lean_Compiler_Simp_inlineApp___lambda__2(x_1, x_2, x_9, x_15, x_3, x_4, x_5, x_6, x_7, x_14); -return x_16; +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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +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_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_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); +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; 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; uint8_t x_31; +x_18 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_ctor_get(x_11, 0); +lean_inc(x_19); lean_dec(x_11); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_20); +x_22 = l_Lean_Compiler_Simp_internalize_visitCases___closed__1; +lean_inc(x_21); +x_23 = lean_mk_array(x_21, x_22); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_21, x_24); +lean_dec(x_21); lean_inc(x_1); -x_18 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_18, 0, x_1); -x_19 = l_Lean_Compiler_Simp_inlineApp___closed__8; -x_20 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7; -x_22 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp___spec__2(x_10, x_22, x_3, x_4, x_5, x_6, x_7, x_17); -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_Lean_Compiler_Simp_inlineApp___lambda__2(x_1, x_2, x_9, x_24, x_3, x_4, x_5, x_6, x_7, x_25); -return x_26; +x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_23, x_25); +x_27 = lean_array_get_size(x_26); +x_28 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__6; +x_29 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(x_28, x_4, x_5, x_6, x_7, x_8, x_18); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_box(0); +x_34 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4(x_19, x_1, x_26, x_27, x_3, x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_32); +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; +x_35 = lean_ctor_get(x_29, 1); +lean_inc(x_35); +lean_dec(x_29); +lean_inc(x_1); +x_36 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_36, 0, x_1); +x_37 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__8; +x_38 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +x_39 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_40 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(x_28, x_40, x_4, x_5, x_6, x_7, x_8, x_35); +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 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4(x_19, x_1, x_26, x_27, x_3, x_2, x_42, x_4, x_5, x_6, x_7, x_8, x_43); +return x_44; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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) { +else +{ +uint8_t x_45; +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_45 = !lean_is_exclusive(x_10); +if (x_45 == 0) +{ +return x_10; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_10, 0); +x_47 = lean_ctor_get(x_10, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_10); +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_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___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_throwError___at_Lean_Compiler_Simp_visitLambda___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_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___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, 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_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_forInUnsafe_loop___at_Lean_Compiler_Simp_visitLambda___spec__4(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___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_8; -x_8 = l_Lean_Compiler_Simp_visitLambda___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_object* x_10; +x_10 = l_Lean_Compiler_Simp_visitLambda___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLambda___boxed(lean_object* x_1, lean_object* x_2, 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); -return x_8; +x_10 = l_Lean_Compiler_Simp_visitLambda(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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) { @@ -8851,6 +12918,16 @@ x_15 = l_Lean_Compiler_Simp_visitLet___lambda__1(x_1, x_2, x_3, x_4, x_14, x_6, return x_15; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_visitLet___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_Compiler_Simp_visitLet___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_4); +lean_dec(x_3); +return x_11; +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_visitCases___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: { @@ -8861,11 +12938,24 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__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_addTrace___at_Lean_Compiler_Simp_inlineApp_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) { _start: { lean_object* x_9; -x_9 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___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); @@ -8874,364 +12964,454 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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_EXPORT lean_object* l_Lean_Compiler_Simp_inlineApp_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, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___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_Lean_Compiler_Simp_inlineApp_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) { _start: { -lean_object* x_8; -x_8 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp___spec__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_object* x_10; +x_10 = l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_3); -lean_dec(x_2); -return x_8; +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp___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_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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, 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_addTrace___at_Lean_Compiler_Simp_inlineApp___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_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_st_ref_get(x_10, x_11); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_get(x_7, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get_uint8(x_15, sizeof(void*)*1); +lean_dec(x_15); +if (x_16 == 0) +{ +uint8_t x_17; lean_dec(x_4); lean_dec(x_3); -return x_9; +lean_dec(x_2); +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_14, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +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; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_Decl_simp_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = lean_st_ref_get(x_3, x_4); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1; -x_8 = lean_st_ref_get(x_7, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_14); +if (x_23 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_2, 2); -x_12 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_10, x_11, x_1); -lean_dec(x_10); -x_13 = lean_box(x_12); -lean_ctor_set(x_8, 0, x_13); -return x_8; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_14, 0); +lean_dec(x_24); +x_25 = lean_alloc_ctor(0, 4, 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); +lean_ctor_set(x_25, 3, x_4); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_14, 0, x_26); +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_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_ctor_get(x_2, 2); -x_17 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_14, x_16, x_1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_14, 1); +lean_inc(x_27); lean_dec(x_14); -x_18 = lean_box(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_15); -return x_19; +x_28 = lean_alloc_ctor(0, 4, 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); +lean_ctor_set(x_28, 3, x_4); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, 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_27); +return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1() { _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; uint8_t x_13; -x_6 = lean_ctor_get(x_3, 5); -x_7 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_3, x_4, x_5); -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_st_ref_take(x_4, x_9); -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_is_exclusive(x_11); -if (x_13 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("stat", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2() { +_start: { -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; uint8_t x_21; -x_14 = lean_ctor_get(x_11, 3); -x_15 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_16 = 0; -x_17 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_8); -lean_ctor_set(x_17, 2, x_15); -lean_ctor_set_uint8(x_17, sizeof(void*)*3, x_16); -lean_inc(x_6); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_6); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_Std_PersistentArray_push___rarg(x_14, x_18); -lean_ctor_set(x_11, 3, x_19); -x_20 = lean_st_ref_set(x_4, x_11, x_12); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(": ", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3() { +_start: { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_20, 0, x_23); -return x_20; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f___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) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_7); +x_14 = l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1; +x_15 = l_Lean_Name_str___override(x_1, x_14); +lean_inc(x_15); +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(x_15, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_15); +lean_dec(x_6); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Compiler_Decl_simp_x3f___lambda__1(x_2, x_3, x_4, x_5, x_20, x_8, x_9, x_10, x_11, x_12, x_19); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_21; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +lean_inc(x_12); +lean_inc(x_11); +x_23 = l_Lean_Compiler_getLCNFSize(x_6, x_11, x_12, 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; 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_23, 0); lean_inc(x_24); -lean_dec(x_20); -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; -} +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_2); +x_26 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_26, 0, x_2); +x_27 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_28 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3; +x_30 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Nat_repr(x_24); +x_32 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_34, 0, x_30); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_27); +x_36 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(x_15, x_35, x_8, x_9, x_10, x_11, x_12, x_25); +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); +x_39 = l_Lean_Compiler_Decl_simp_x3f___lambda__1(x_2, x_3, x_4, x_5, x_37, x_8, x_9, x_10, x_11, x_12, x_38); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_37); +return x_39; } 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; 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; -x_27 = lean_ctor_get(x_11, 0); -x_28 = lean_ctor_get(x_11, 1); -x_29 = lean_ctor_get(x_11, 2); -x_30 = lean_ctor_get(x_11, 3); -x_31 = lean_ctor_get(x_11, 4); -x_32 = lean_ctor_get(x_11, 5); -x_33 = lean_ctor_get(x_11, 6); -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); +uint8_t x_40; +lean_dec(x_15); +lean_dec(x_12); lean_dec(x_11); -x_34 = l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7; -x_35 = 0; -x_36 = lean_alloc_ctor(12, 3, 1); -lean_ctor_set(x_36, 0, x_1); -lean_ctor_set(x_36, 1, x_8); -lean_ctor_set(x_36, 2, x_34); -lean_ctor_set_uint8(x_36, sizeof(void*)*3, x_35); -lean_inc(x_6); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_6); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Std_PersistentArray_push___rarg(x_30, x_37); -x_39 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_39, 0, x_27); -lean_ctor_set(x_39, 1, x_28); -lean_ctor_set(x_39, 2, x_29); -lean_ctor_set(x_39, 3, x_38); -lean_ctor_set(x_39, 4, x_31); -lean_ctor_set(x_39, 5, x_32); -lean_ctor_set(x_39, 6, x_33); -x_40 = lean_st_ref_set(x_4, x_39, x_12); -x_41 = lean_ctor_get(x_40, 1); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_23); +if (x_40 == 0) +{ +return x_23; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_23, 0); +x_42 = lean_ctor_get(x_23, 1); +lean_inc(x_42); 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); +lean_dec(x_23); +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; } -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; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("new", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" :=\n", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f___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_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; uint8_t 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_unsigned_to_nat(2u); -x_11 = 1; -x_12 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_1); -lean_ctor_set_uint8(x_12, sizeof(void*)*2, x_11); -x_13 = lean_st_ref_get(x_8, x_9); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Compiler_Simp_collectInlineStats___closed__7; -x_16 = lean_st_mk_ref(x_15, x_14); +uint8_t x_15; lean_object* x_16; +lean_dec(x_8); +x_15 = 0; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_16 = l_Lean_Compiler_Simp_visitLambda(x_1, x_15, x_9, x_10, x_11, x_12, x_13, 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; uint8_t x_23; 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_get(x_8, x_18); -x_20 = lean_ctor_get(x_19, 1); +x_19 = l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1; +x_20 = l_Lean_Name_str___override(x_2, x_19); lean_inc(x_20); -lean_dec(x_19); -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_st_mk_ref(x_22, x_20); -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_8); -lean_inc(x_17); +x_21 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(x_20, x_9, x_10, x_11, x_12, x_13, x_18); +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) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_20); +x_24 = lean_ctor_get(x_21, 1); lean_inc(x_24); -x_26 = l_Lean_Compiler_Simp_visitLambda(x_2, x_12, x_24, x_17, x_7, x_8, x_25); -if (lean_obj_tag(x_26) == 0) +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = l_Lean_Compiler_Decl_simp_x3f___lambda__2(x_3, x_4, x_5, x_6, x_17, x_7, x_25, x_9, x_10, x_11, x_12, x_13, x_24); +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; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_27 = lean_ctor_get(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; +x_27 = lean_ctor_get(x_21, 1); lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_st_ref_get(x_8, x_28); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_st_ref_get(x_24, x_30); -lean_dec(x_24); -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_st_ref_get(x_8, x_33); -lean_dec(x_8); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -x_36 = lean_st_ref_get(x_17, x_35); -lean_dec(x_17); -x_37 = lean_unbox(x_32); -lean_dec(x_32); -if (x_37 == 0) +lean_dec(x_21); +lean_inc(x_4); +x_28 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_28, 0, x_4); +x_29 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_30 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3; +x_32 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_inc(x_17); +x_33 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_33, 0, x_17); +x_34 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_29); +x_36 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(x_20, x_35, x_9, x_10, x_11, x_12, x_13, x_27); +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); +x_39 = l_Lean_Compiler_Decl_simp_x3f___lambda__2(x_3, x_4, x_5, x_6, x_17, x_7, x_37, x_9, x_10, x_11, x_12, x_13, x_38); +return x_39; +} +} +else { -uint8_t x_38; -lean_dec(x_27); +uint8_t x_40; +lean_dec(x_13); +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_4); lean_dec(x_3); -x_38 = !lean_is_exclusive(x_36); -if (x_38 == 0) +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_16); +if (x_40 == 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; +return x_16; } else { lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_36, 1); +x_41 = lean_ctor_get(x_16, 0); +x_42 = lean_ctor_get(x_16, 1); +lean_inc(x_42); 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); +lean_dec(x_16); +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 -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_36); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_36, 0); -lean_dec(x_45); -x_46 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_46, 0, x_3); -lean_ctor_set(x_46, 1, x_4); -lean_ctor_set(x_46, 2, x_5); -lean_ctor_set(x_46, 3, x_27); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_36, 0, x_47); -return x_36; } -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_36, 1); -lean_inc(x_48); -lean_dec(x_36); -x_49 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_49, 0, x_3); -lean_ctor_set(x_49, 1, x_4); -lean_ctor_set(x_49, 2, x_5); -lean_ctor_set(x_49, 3, x_27); -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_48); -return x_51; } +static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("step", 4); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f___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) { +_start: { -uint8_t x_52; -lean_dec(x_24); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_7); +x_14 = l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1; +lean_inc(x_1); +x_15 = l_Lean_Name_str___override(x_1, x_14); +lean_inc(x_15); +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(x_15, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); lean_dec(x_17); -lean_dec(x_8); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_52 = !lean_is_exclusive(x_26); -if (x_52 == 0) +if (x_18 == 0) { -return x_26; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Compiler_Decl_simp_x3f___lambda__3(x_2, x_15, x_1, x_3, x_4, x_5, x_6, x_20, x_8, x_9, x_10, x_11, x_12, x_19); +return x_21; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_26, 0); -x_54 = lean_ctor_get(x_26, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_26); -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_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_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +lean_inc(x_3); +x_23 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_23, 0, x_3); +x_24 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_25 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3; +x_27 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_inc(x_6); +x_28 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_28, 0, x_6); +x_29 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +lean_inc(x_15); +x_31 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(x_15, x_30, x_8, x_9, x_10, x_11, x_12, x_22); +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_Lean_Compiler_Decl_simp_x3f___lambda__3(x_2, x_15, x_1, x_3, x_4, x_5, x_6, x_32, x_8, x_9, x_10, x_11, x_12, x_33); +return x_34; } } } @@ -9239,7 +13419,7 @@ static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("stats", 5); +x_1 = lean_mk_string_from_bytes("occs", 4); return x_1; } } @@ -9247,7 +13427,7 @@ static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_inlineApp___closed__6; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__6; x_2 = l_Lean_Compiler_Decl_simp_x3f___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -9270,43 +13450,26 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Compiler_Decl_simp_x3f___closed__5() { -_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; -} -} -LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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_5; -lean_inc(x_3); -lean_inc(x_2); -x_5 = l_Lean_Compiler_Decl_ensureUniqueLetVarNames(x_1, 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); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_ctor_get(x_6, 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); lean_inc(x_8); -x_9 = lean_ctor_get(x_6, 1); +x_9 = lean_ctor_get(x_1, 1); lean_inc(x_9); -x_10 = lean_ctor_get(x_6, 2); +x_10 = lean_ctor_get(x_1, 2); lean_inc(x_10); -x_11 = lean_ctor_get(x_6, 3); +x_11 = lean_ctor_get(x_1, 3); lean_inc(x_11); -lean_dec(x_6); +lean_dec(x_1); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_11); -x_12 = l_Lean_Compiler_Simp_collectInlineStats(x_11, x_2, x_3, x_7); +x_12 = l_Lean_Compiler_Simp_internalize_visitLambda(x_11, x_2, x_3, x_4, x_5, x_6, 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; @@ -9316,251 +13479,358 @@ x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); x_15 = l_Lean_Compiler_Decl_simp_x3f___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Decl_simp_x3f___spec__1(x_15, x_2, x_3, x_14); +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1(x_15, x_2, x_3, x_4, x_5, x_6, x_14); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_unbox(x_17); lean_dec(x_17); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); lean_dec(x_16); -x_20 = lean_box(0); -x_21 = l_Lean_Compiler_Decl_simp_x3f___lambda__1(x_13, x_11, x_8, x_9, x_10, x_20, x_2, x_3, x_19); -return x_21; +x_20 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +x_21 = lean_box(0); +x_22 = l_Lean_Compiler_Decl_simp_x3f___lambda__4(x_20, x_13, x_8, x_9, x_10, x_11, x_21, x_2, x_3, x_4, x_5, x_6, x_19); +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; 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_22 = lean_ctor_get(x_16, 1); -lean_inc(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; +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); lean_dec(x_16); +x_24 = lean_st_ref_get(x_6, x_23); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_st_ref_get(x_3, x_25); +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_inc(x_8); -x_23 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_23, 0, x_8); -x_24 = l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7; -x_25 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_Compiler_Decl_simp_x3f___closed__4; -x_27 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -lean_inc(x_13); -x_28 = l_Lean_Compiler_Simp_InlineStats_format(x_13); -x_29 = l_Lean_Compiler_Decl_simp_x3f___closed__5; -x_30 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = lean_alloc_ctor(0, 1, 0); +x_29 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_29, 0, x_8); +x_30 = l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3; +x_31 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_Compiler_Decl_simp_x3f___closed__4; x_33 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_24); -x_34 = l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2(x_15, x_33, x_2, x_3, x_22); -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 = l_Lean_Compiler_Decl_simp_x3f___lambda__1(x_13, x_11, x_8, x_9, x_10, x_35, x_2, x_3, x_36); -lean_dec(x_35); -return x_37; +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_ctor_get(x_27, 0); +lean_inc(x_34); +lean_dec(x_27); +x_35 = l_Lean_Compiler_Simp_OccInfo_format(x_34); +x_36 = l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3; +x_37 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_39, 0, x_33); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_30); +x_41 = l_Lean_addTrace___at_Lean_Compiler_Simp_inlineApp_x3f___spec__2(x_15, x_40, x_2, x_3, x_4, x_5, x_6, x_28); +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 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +x_45 = l_Lean_Compiler_Decl_simp_x3f___lambda__4(x_44, x_13, x_8, x_9, x_10, x_11, x_42, x_2, x_3, x_4, x_5, x_6, x_43); +return x_45; } } else { -uint8_t x_38; +uint8_t x_46; 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_4); lean_dec(x_3); lean_dec(x_2); -x_38 = !lean_is_exclusive(x_12); -if (x_38 == 0) +x_46 = !lean_is_exclusive(x_12); +if (x_46 == 0) { return x_12; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_12, 0); -x_40 = lean_ctor_get(x_12, 1); -lean_inc(x_40); -lean_inc(x_39); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_12, 0); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_12); -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; +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 +LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_simp_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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_42; -lean_dec(x_3); -lean_dec(x_2); -x_42 = !lean_is_exclusive(x_5); -if (x_42 == 0) +lean_object* x_12; +x_12 = l_Lean_Compiler_Decl_simp_x3f___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_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; +} +} +static lean_object* _init_l_Lean_Compiler_Decl_simp___closed__1() { +_start: { -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2; +x_2 = lean_unsigned_to_nat(0u); +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 +} +static lean_object* _init_l_Lean_Compiler_Decl_simp___closed__2() { +_start: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +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_Compiler_Decl_simp___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_Decl_simp___closed__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_Lean_isTracingEnabledFor___at_Lean_Compiler_Decl_simp_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Compiler_Decl_simp___closed__4() { _start: { -lean_object* x_5; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_Decl_simp_x3f___spec__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); +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_Compiler_Decl_simp___closed__3; +x_3 = l_Lean_Compiler_Decl_simp___closed__2; +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; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Compiler_Decl_simp___closed__5() { _start: { -lean_object* x_6; -x_6 = l_Lean_addTrace___at_Lean_Compiler_Decl_simp_x3f___spec__2(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_Decl_simp___closed__1; +x_2 = l_Lean_Compiler_Decl_simp___closed__4; +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_Compiler_Decl_simp_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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Compiler_Decl_simp___closed__6() { _start: { -lean_object* x_10; -x_10 = l_Lean_Compiler_Decl_simp_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_6); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Compiler_Decl_simp___closed__5; +x_2 = l_Lean_Compiler_Simp_internalize_visitLambda___closed__1; +x_3 = lean_unsigned_to_nat(1u); +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_Lean_Compiler_Decl_simp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_st_ref_get(x_3, x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Compiler_Decl_simp___closed__6; +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 = lean_st_ref_get(x_3, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Compiler_Simp_instInhabitedState___closed__1; +x_14 = lean_st_mk_ref(x_13, x_12); +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_unsigned_to_nat(1u); lean_inc(x_3); lean_inc(x_2); +lean_inc(x_9); +lean_inc(x_15); lean_inc(x_1); -x_5 = l_Lean_Compiler_Decl_simp_x3f(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) +x_18 = l_Lean_Compiler_Decl_simp_x3f(x_1, x_17, x_15, x_9, x_2, x_3, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_6; -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -if (lean_obj_tag(x_6) == 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; +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_st_ref_get(x_3, x_20); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_get(x_15, x_22); +lean_dec(x_15); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_get(x_3, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_9, x_26); +lean_dec(x_9); +if (lean_obj_tag(x_19) == 0) { -uint8_t x_7; +uint8_t x_28; lean_dec(x_3); lean_dec(x_2); -x_7 = !lean_is_exclusive(x_5); -if (x_7 == 0) +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_8; -x_8 = lean_ctor_get(x_5, 0); -lean_dec(x_8); -lean_ctor_set(x_5, 0, x_1); -return x_5; +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_1); +return x_27; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_5, 1); -lean_inc(x_9); -lean_dec(x_5); -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_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_1); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -lean_object* x_11; lean_object* x_12; +lean_object* x_32; lean_object* x_33; lean_dec(x_1); -x_11 = lean_ctor_get(x_5, 1); -lean_inc(x_11); -lean_dec(x_5); -x_12 = lean_ctor_get(x_6, 0); -lean_inc(x_12); -lean_dec(x_6); -x_1 = x_12; -x_4 = x_11; +x_32 = lean_ctor_get(x_27, 1); +lean_inc(x_32); +lean_dec(x_27); +x_33 = lean_ctor_get(x_19, 0); +lean_inc(x_33); +lean_dec(x_19); +x_1 = x_33; +x_4 = x_32; goto _start; } } else { -uint8_t x_14; +uint8_t x_35; +lean_dec(x_15); +lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = !lean_is_exclusive(x_5); -if (x_14 == 0) +x_35 = !lean_is_exclusive(x_18); +if (x_35 == 0) { -return x_5; +return x_18; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_5, 0); -x_16 = lean_ctor_get(x_5, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_5); -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_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; } } } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("step", 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +x_2 = l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__4; +x_2 = l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2() { +static lean_object* _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_Simp_inlineApp___closed__4; -x_2 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1; +x_1 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2; +x_2 = l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; -x_2 = l_Lean_Compiler_Simp_inlineApp___closed__6; +x_2 = l_Lean_Compiler_Simp_inlineApp_x3f___closed__6; x_3 = 0; x_4 = l_Lean_registerTraceClass(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) @@ -9569,7 +13839,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); -x_6 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2; +x_6 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1; x_7 = l_Lean_registerTraceClass(x_6, x_3, x_5); if (lean_obj_tag(x_7) == 0) { @@ -9577,53 +13847,115 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); -x_9 = l_Lean_Compiler_Decl_simp_x3f___closed__2; +x_9 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2; x_10 = l_Lean_registerTraceClass(x_9, x_3, x_8); +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, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3; +x_13 = l_Lean_registerTraceClass(x_12, x_3, x_11); +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, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_Compiler_Decl_simp_x3f___closed__2; +x_16 = l_Lean_registerTraceClass(x_15, x_3, x_14); +return x_16; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) +{ +return x_13; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +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 = 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 +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) +{ return x_10; } else { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_7); -if (x_11 == 0) +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; +} +} +} +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_7); +if (x_25 == 0) { return x_7; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_7, 0); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_7, 0); +x_27 = lean_ctor_get(x_7, 1); +lean_inc(x_27); +lean_inc(x_26); lean_dec(x_7); -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; +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 { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_4); -if (x_15 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_4); +if (x_29 == 0) { return x_4; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_4, 0); +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); +lean_inc(x_30); lean_dec(x_4); -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; +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; } } } @@ -9653,106 +13985,134 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_InlineAttrs(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1 = _init_l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_InlineStats_numOccs___default___closed__1); -l_Lean_Compiler_Simp_InlineStats_numOccs___default = _init_l_Lean_Compiler_Simp_InlineStats_numOccs___default(); -lean_mark_persistent(l_Lean_Compiler_Simp_InlineStats_numOccs___default); -l_Lean_Compiler_Simp_InlineStats_size___default = _init_l_Lean_Compiler_Simp_InlineStats_size___default(); -lean_mark_persistent(l_Lean_Compiler_Simp_InlineStats_size___default); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__1); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__2); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__3); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__4); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__5); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__6); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__7); -l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_InlineStats_format___spec__4___closed__8); -l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1 = _init_l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_instToFormatInlineStats___closed__1); -l_Lean_Compiler_Simp_instToFormatInlineStats = _init_l_Lean_Compiler_Simp_instToFormatInlineStats(); -lean_mark_persistent(l_Lean_Compiler_Simp_instToFormatInlineStats); -l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1 = _init_l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goLambda___closed__1); -l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1(); -lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__1); -l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2(); -lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__2); -l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__3 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__3(); -lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__3); -l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4(); -lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_Simp_collectInlineStats_go___spec__1___closed__4); -l_Lean_Compiler_Simp_collectInlineStats_go___closed__1 = _init_l_Lean_Compiler_Simp_collectInlineStats_go___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_go___closed__1); -l_Lean_Compiler_Simp_collectInlineStats_go___closed__2 = _init_l_Lean_Compiler_Simp_collectInlineStats_go___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_go___closed__2); -l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1(); -lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__1___closed__1); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__1); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__2); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__3); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__4); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__5); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__6); -l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7 = _init_l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Compiler_Simp_collectInlineStats_goValue___spec__2___closed__7); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__1 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__1); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__2); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__3); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__4); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__5); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__6); -l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7 = _init_l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats_goValue___closed__7); -l_Lean_Compiler_Simp_collectInlineStats___closed__1 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__1); -l_Lean_Compiler_Simp_collectInlineStats___closed__2 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__2); -l_Lean_Compiler_Simp_collectInlineStats___closed__3 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__3); -l_Lean_Compiler_Simp_collectInlineStats___closed__4 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__4(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__4); -l_Lean_Compiler_Simp_collectInlineStats___closed__5 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__5(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__5); -l_Lean_Compiler_Simp_collectInlineStats___closed__6 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__6(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__6); -l_Lean_Compiler_Simp_collectInlineStats___closed__7 = _init_l_Lean_Compiler_Simp_collectInlineStats___closed__7(); -lean_mark_persistent(l_Lean_Compiler_Simp_collectInlineStats___closed__7); -l_Lean_Compiler_Simp_Config_increaseFactor___default = _init_l_Lean_Compiler_Simp_Config_increaseFactor___default(); -lean_mark_persistent(l_Lean_Compiler_Simp_Config_increaseFactor___default); +l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1 = _init_l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_Occ_noConfusion___rarg___closed__1); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__1 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__1(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__1); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__2); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__3); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__4); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__5 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__5(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__5); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__6); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__7); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__8 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__8(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__8); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__9); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__10); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__11); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__12 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__12(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__12); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__13); +l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14 = _init_l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14(); +lean_mark_persistent(l___private_Lean_Compiler_Simp_0__Lean_Compiler_Simp_reprOcc____x40_Lean_Compiler_Simp___hyg_296____closed__14); +l_Lean_Compiler_Simp_instReprOcc___closed__1 = _init_l_Lean_Compiler_Simp_instReprOcc___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_instReprOcc___closed__1); +l_Lean_Compiler_Simp_instReprOcc = _init_l_Lean_Compiler_Simp_instReprOcc(); +lean_mark_persistent(l_Lean_Compiler_Simp_instReprOcc); +l_Lean_Compiler_Simp_instInhabitedOcc = _init_l_Lean_Compiler_Simp_instInhabitedOcc(); +l_Lean_Compiler_Simp_OccInfo_map___default = _init_l_Lean_Compiler_Simp_OccInfo_map___default(); +lean_mark_persistent(l_Lean_Compiler_Simp_OccInfo_map___default); +l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1 = _init_l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_instInhabitedOccInfo___closed__1); +l_Lean_Compiler_Simp_instInhabitedOccInfo = _init_l_Lean_Compiler_Simp_instInhabitedOccInfo(); +lean_mark_persistent(l_Lean_Compiler_Simp_instInhabitedOccInfo); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__1 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__1); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__2); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__3); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__4); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__5); +l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6 = _init_l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Compiler_Simp_OccInfo_format___spec__4___closed__6); +l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1 = _init_l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_instToFormatOccInfo___closed__1); +l_Lean_Compiler_Simp_instToFormatOccInfo = _init_l_Lean_Compiler_Simp_instToFormatOccInfo(); +lean_mark_persistent(l_Lean_Compiler_Simp_instToFormatOccInfo); +l_Lean_Compiler_Simp_Config_smallThreshold___default = _init_l_Lean_Compiler_Simp_Config_smallThreshold___default(); +lean_mark_persistent(l_Lean_Compiler_Simp_Config_smallThreshold___default); l_Lean_Compiler_Simp_Context_config___default = _init_l_Lean_Compiler_Simp_Context_config___default(); lean_mark_persistent(l_Lean_Compiler_Simp_Context_config___default); -l_Lean_Compiler_Simp_Context_localInline___default = _init_l_Lean_Compiler_Simp_Context_localInline___default(); +l_Lean_Compiler_Simp_State_occInfo___default = _init_l_Lean_Compiler_Simp_State_occInfo___default(); +lean_mark_persistent(l_Lean_Compiler_Simp_State_occInfo___default); l_Lean_Compiler_Simp_State_simplified___default = _init_l_Lean_Compiler_Simp_State_simplified___default(); +l_Lean_Compiler_Simp_instInhabitedState___closed__1 = _init_l_Lean_Compiler_Simp_instInhabitedState___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_instInhabitedState___closed__1); +l_Lean_Compiler_Simp_instInhabitedState = _init_l_Lean_Compiler_Simp_instInhabitedState(); +lean_mark_persistent(l_Lean_Compiler_Simp_instInhabitedState); +l_Lean_Compiler_Simp_internalize_visitLambda___closed__1 = _init_l_Lean_Compiler_Simp_internalize_visitLambda___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_internalize_visitLambda___closed__1); +l_Lean_Compiler_Simp_internalize_visitCases___closed__1 = _init_l_Lean_Compiler_Simp_internalize_visitCases___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_internalize_visitCases___closed__1); +l_Lean_Compiler_Simp_simpProj_x3f___closed__1 = _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_simpProj_x3f___closed__1); +l_Lean_Compiler_Simp_simpProj_x3f___closed__2 = _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Simp_simpProj_x3f___closed__2); +l_Lean_Compiler_Simp_simpProj_x3f___closed__3 = _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Simp_simpProj_x3f___closed__3); +l_Lean_Compiler_Simp_simpProj_x3f___closed__4 = _init_l_Lean_Compiler_Simp_simpProj_x3f___closed__4(); +lean_mark_persistent(l_Lean_Compiler_Simp_simpProj_x3f___closed__4); +l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1 = _init_l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_simpAppApp_x3f___closed__1); +l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1 = _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__1); +l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2 = _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__2); +l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3 = _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3(); +lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__3); +l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4 = _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4(); +lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__4); +l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5 = _init_l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5(); +lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___spec__1___closed__5); +l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1 = _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__1); +l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2 = _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__2); +l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3 = _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__3); +l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4 = _init_l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineProjInst_x3f_visitProj___closed__4); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__1 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__1(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__1); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__2); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__3); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__4); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__5); +l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6 = _init_l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6(); +lean_mark_persistent(l_Lean_throwError___at_Lean_Compiler_Simp_visitLambda___spec__3___closed__6); +l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1 = _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1(); +lean_mark_persistent(l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__1); +l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2 = _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2(); +lean_mark_persistent(l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__2); +l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3 = _init_l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3(); +lean_mark_persistent(l_Lean_Compiler_getLevel___at_Lean_Compiler_Simp_visitLambda___spec__2___closed__3); +l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1 = _init_l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__1); +l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2 = _init_l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Compiler_mkLcUnreachable___at_Lean_Compiler_Simp_visitLambda___spec__1___closed__2); l_Lean_Compiler_Simp_visitLambda___closed__1 = _init_l_Lean_Compiler_Simp_visitLambda___closed__1(); lean_mark_persistent(l_Lean_Compiler_Simp_visitLambda___closed__1); l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1 = _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1(); lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__1); l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2 = _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2(); lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__2); -l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3 = _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3(); -lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__3); -l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4 = _init_l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4(); -lean_mark_persistent(l_panic___at_Lean_Compiler_Simp_visitCases___spec__2___closed__4); l_Lean_Compiler_Simp_visitCases___closed__1 = _init_l_Lean_Compiler_Simp_visitCases___closed__1(); lean_mark_persistent(l_Lean_Compiler_Simp_visitCases___closed__1); l_Lean_Compiler_Simp_visitCases___closed__2 = _init_l_Lean_Compiler_Simp_visitCases___closed__2(); @@ -9763,56 +14123,58 @@ l_Lean_Compiler_Simp_visitCases___closed__4 = _init_l_Lean_Compiler_Simp_visitCa lean_mark_persistent(l_Lean_Compiler_Simp_visitCases___closed__4); l_Lean_Compiler_Simp_visitCases___closed__5 = _init_l_Lean_Compiler_Simp_visitCases___closed__5(); lean_mark_persistent(l_Lean_Compiler_Simp_visitCases___closed__5); -l_Lean_Compiler_Simp_visitCases___closed__6 = _init_l_Lean_Compiler_Simp_visitCases___closed__6(); -lean_mark_persistent(l_Lean_Compiler_Simp_visitCases___closed__6); -l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__2___closed__1); +l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Compiler_Simp_inlineApp_x3f___spec__1___closed__1); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__1___closed__1); l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__1); l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__2); -l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__3); -l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__4); -l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__5); -l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__3___closed__6); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__1); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__2); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__3); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__4); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__5); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__6); +l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___lambda__4___closed__7); l_Lean_Compiler_Simp_inlineApp_x3f___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__1(); lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__1); l_Lean_Compiler_Simp_inlineApp_x3f___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__2(); lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__2); -l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__1); -l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__2); -l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__3); -l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__1___closed__4); -l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__1); -l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__2); -l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___lambda__2___closed__3); -l_Lean_Compiler_Simp_inlineApp___closed__1 = _init_l_Lean_Compiler_Simp_inlineApp___closed__1(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__1); -l_Lean_Compiler_Simp_inlineApp___closed__2 = _init_l_Lean_Compiler_Simp_inlineApp___closed__2(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__2); -l_Lean_Compiler_Simp_inlineApp___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp___closed__3(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__3); -l_Lean_Compiler_Simp_inlineApp___closed__4 = _init_l_Lean_Compiler_Simp_inlineApp___closed__4(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__4); -l_Lean_Compiler_Simp_inlineApp___closed__5 = _init_l_Lean_Compiler_Simp_inlineApp___closed__5(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__5); -l_Lean_Compiler_Simp_inlineApp___closed__6 = _init_l_Lean_Compiler_Simp_inlineApp___closed__6(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__6); -l_Lean_Compiler_Simp_inlineApp___closed__7 = _init_l_Lean_Compiler_Simp_inlineApp___closed__7(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__7); -l_Lean_Compiler_Simp_inlineApp___closed__8 = _init_l_Lean_Compiler_Simp_inlineApp___closed__8(); -lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp___closed__8); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__3 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__3); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__4 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__4(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__4); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__5 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__5(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__5); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__6 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__6(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__6); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__7 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__7(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__7); +l_Lean_Compiler_Simp_inlineApp_x3f___closed__8 = _init_l_Lean_Compiler_Simp_inlineApp_x3f___closed__8(); +lean_mark_persistent(l_Lean_Compiler_Simp_inlineApp_x3f___closed__8); +l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__1); +l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__2); +l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__2___closed__3); +l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__1); +l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__2); +l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__3___closed__3); +l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1 = _init_l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___lambda__4___closed__1); l_Lean_Compiler_Decl_simp_x3f___closed__1 = _init_l_Lean_Compiler_Decl_simp_x3f___closed__1(); lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___closed__1); l_Lean_Compiler_Decl_simp_x3f___closed__2 = _init_l_Lean_Compiler_Decl_simp_x3f___closed__2(); @@ -9821,13 +14183,25 @@ l_Lean_Compiler_Decl_simp_x3f___closed__3 = _init_l_Lean_Compiler_Decl_simp_x3f_ lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___closed__3); l_Lean_Compiler_Decl_simp_x3f___closed__4 = _init_l_Lean_Compiler_Decl_simp_x3f___closed__4(); lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___closed__4); -l_Lean_Compiler_Decl_simp_x3f___closed__5 = _init_l_Lean_Compiler_Decl_simp_x3f___closed__5(); -lean_mark_persistent(l_Lean_Compiler_Decl_simp_x3f___closed__5); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__1); -l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2(); -lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710____closed__2); -res = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_3710_(lean_io_mk_world()); +l_Lean_Compiler_Decl_simp___closed__1 = _init_l_Lean_Compiler_Decl_simp___closed__1(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__1); +l_Lean_Compiler_Decl_simp___closed__2 = _init_l_Lean_Compiler_Decl_simp___closed__2(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__2); +l_Lean_Compiler_Decl_simp___closed__3 = _init_l_Lean_Compiler_Decl_simp___closed__3(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__3); +l_Lean_Compiler_Decl_simp___closed__4 = _init_l_Lean_Compiler_Decl_simp___closed__4(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__4); +l_Lean_Compiler_Decl_simp___closed__5 = _init_l_Lean_Compiler_Decl_simp___closed__5(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__5); +l_Lean_Compiler_Decl_simp___closed__6 = _init_l_Lean_Compiler_Decl_simp___closed__6(); +lean_mark_persistent(l_Lean_Compiler_Decl_simp___closed__6); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__1); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__2); +l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3 = _init_l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3(); +lean_mark_persistent(l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456____closed__3); +res = l_Lean_Compiler_initFn____x40_Lean_Compiler_Simp___hyg_6456_(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/Stage1.c b/stage0/stdlib/Lean/Compiler/Stage1.c index b39246e4dbdc..952a796dae25 100644 --- a/stage0/stdlib/Lean/Compiler/Stage1.c +++ b/stage0/stdlib/Lean/Compiler/Stage1.c @@ -54,6 +54,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_compileStage1___boxed(lean_object*, lea static size_t l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_saveStage1Decls___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_getStage1Decl_x3f___lambda__1(lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Compiler_getStage1Decl_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Compiler_saveStage1Decls___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1207,342 +1208,505 @@ lean_inc(x_1); x_14 = l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(x_1, x_2, x_3, x_8); if (lean_obj_tag(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; -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_Lean_ConstantInfo_all(x_15); -lean_dec(x_15); -x_18 = l_List_redLength___rarg(x_17); -x_19 = lean_mk_empty_array_with_capacity(x_18); -lean_dec(x_18); -x_20 = l_List_toArrayAux___rarg(x_17, x_19); -x_21 = lean_compile_stage1(x_20, x_2, x_3, x_16); -if (lean_obj_tag(x_21) == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = l_Lean_ConstantInfo_hasValue(x_16); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_16); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_19 = lean_box(0); +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; lean_object* x_24; +lean_free_object(x_14); +x_20 = l_Lean_ConstantInfo_all(x_16); +lean_dec(x_16); +x_21 = l_List_redLength___rarg(x_20); +x_22 = lean_mk_empty_array_with_capacity(x_21); +lean_dec(x_21); +x_23 = l_List_toArrayAux___rarg(x_20, x_22); +x_24 = lean_compile_stage1(x_23, x_2, x_3, x_17); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +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; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_array_get_size(x_26); +x_28 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_29 = 0; +x_30 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; +x_31 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_30, x_26, x_28, x_29, x_30); +lean_dec(x_26); +lean_dec(x_1); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +lean_dec(x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +x_33 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; +lean_ctor_set(x_24, 0, x_33); +return x_24; +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_dec(x_32); +lean_ctor_set(x_24, 0, x_34); +return x_24; +} +} +else { -lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_array_get_size(x_23); -x_25 = lean_usize_of_nat(x_24); +lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_35 = lean_ctor_get(x_24, 0); +x_36 = lean_ctor_get(x_24, 1); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_24); -x_26 = 0; -x_27 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; -x_28 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_27, x_23, x_25, x_26, x_27); -lean_dec(x_23); +x_37 = lean_array_get_size(x_35); +x_38 = lean_usize_of_nat(x_37); +lean_dec(x_37); +x_39 = 0; +x_40 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; +x_41 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_40, x_35, x_38, x_39, x_40); +lean_dec(x_35); lean_dec(x_1); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -lean_dec(x_28); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; -x_30 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; -lean_ctor_set(x_21, 0, x_30); -return x_21; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_36); +return x_44; } else { -lean_object* x_31; -x_31 = lean_ctor_get(x_29, 0); -lean_inc(x_31); -lean_dec(x_29); -lean_ctor_set(x_21, 0, x_31); -return x_21; +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 0); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; +} } } else { -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_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_21); -x_34 = lean_array_get_size(x_32); -x_35 = lean_usize_of_nat(x_34); -lean_dec(x_34); -x_36 = 0; -x_37 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; -x_38 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_37, x_32, x_35, x_36, x_37); -lean_dec(x_32); +uint8_t x_47; lean_dec(x_1); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -lean_dec(x_38); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; -x_40 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_33); -return x_41; +x_47 = !lean_is_exclusive(x_24); +if (x_47 == 0) +{ +return x_24; } else { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_39, 0); -lean_inc(x_42); -lean_dec(x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_33); -return x_43; +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_24, 0); +x_49 = lean_ctor_get(x_24, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_24); +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 { -uint8_t x_44; +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = lean_ctor_get(x_14, 0); +x_52 = lean_ctor_get(x_14, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_14); +x_53 = l_Lean_ConstantInfo_hasValue(x_51); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_51); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_21); -if (x_44 == 0) +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_52); +return x_55; +} +else { -return x_21; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_56 = l_Lean_ConstantInfo_all(x_51); +lean_dec(x_51); +x_57 = l_List_redLength___rarg(x_56); +x_58 = lean_mk_empty_array_with_capacity(x_57); +lean_dec(x_57); +x_59 = l_List_toArrayAux___rarg(x_56, x_58); +x_60 = lean_compile_stage1(x_59, x_2, x_3, x_52); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_63 = x_60; +} else { + lean_dec_ref(x_60); + x_63 = lean_box(0); +} +x_64 = lean_array_get_size(x_61); +x_65 = lean_usize_of_nat(x_64); +lean_dec(x_64); +x_66 = 0; +x_67 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; +x_68 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_67, x_61, x_65, x_66, x_67); +lean_dec(x_61); +lean_dec(x_1); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +lean_dec(x_68); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; +x_70 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; +if (lean_is_scalar(x_63)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_63; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_62); +return x_71; } 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_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_object* x_72; lean_object* x_73; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +lean_dec(x_69); +if (lean_is_scalar(x_63)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = x_63; +} +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_62); +return x_73; +} +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_1); +x_74 = lean_ctor_get(x_60, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_60, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_76 = x_60; +} else { + lean_dec_ref(x_60); + x_76 = lean_box(0); +} +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(1, 2, 0); +} else { + x_77 = x_76; +} +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_75); +return x_77; +} } } } else { -uint8_t x_48; +uint8_t x_78; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_14); -if (x_48 == 0) +x_78 = !lean_is_exclusive(x_14); +if (x_78 == 0) { return x_14; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_14, 0); -x_50 = lean_ctor_get(x_14, 1); -lean_inc(x_50); -lean_inc(x_49); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_14, 0); +x_80 = lean_ctor_get(x_14, 1); +lean_inc(x_80); +lean_inc(x_79); lean_dec(x_14); -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; +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; } } } else { -uint8_t x_52; +uint8_t x_82; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_52 = !lean_is_exclusive(x_13); -if (x_52 == 0) +x_82 = !lean_is_exclusive(x_13); +if (x_82 == 0) { lean_ctor_set(x_5, 0, x_13); return x_5; } else { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_13, 0); -lean_inc(x_53); +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_13, 0); +lean_inc(x_83); lean_dec(x_13); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_5, 0, x_54); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_5, 0, x_84); return x_5; } } } 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; -x_55 = lean_ctor_get(x_5, 0); -x_56 = lean_ctor_get(x_5, 1); -lean_inc(x_56); -lean_inc(x_55); +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_5, 0); +x_86 = lean_ctor_get(x_5, 1); +lean_inc(x_86); +lean_inc(x_85); lean_dec(x_5); -x_57 = lean_ctor_get(x_55, 0); -lean_inc(x_57); -lean_dec(x_55); -x_58 = l_Lean_Compiler_instInhabitedStage1ExtState; -x_59 = l_Lean_Compiler_saveStage1Decls___closed__1; -x_60 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_58, x_59, x_57); -lean_dec(x_57); +x_87 = lean_ctor_get(x_85, 0); +lean_inc(x_87); +lean_dec(x_85); +x_88 = l_Lean_Compiler_instInhabitedStage1ExtState; +x_89 = l_Lean_Compiler_saveStage1Decls___closed__1; +x_90 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_88, x_89, x_87); +lean_dec(x_87); lean_inc(x_1); -x_61 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_getStage1Decl_x3f___spec__1(x_60, x_1); -if (lean_obj_tag(x_61) == 0) +x_91 = l_Std_PersistentHashMap_find_x3f___at_Lean_Compiler_getStage1Decl_x3f___spec__1(x_90, x_1); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_62; +lean_object* x_92; lean_inc(x_1); -x_62 = l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(x_1, x_2, x_3, x_56); -if (lean_obj_tag(x_62) == 0) -{ -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_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = l_Lean_ConstantInfo_all(x_63); -lean_dec(x_63); -x_66 = l_List_redLength___rarg(x_65); -x_67 = lean_mk_empty_array_with_capacity(x_66); -lean_dec(x_66); -x_68 = l_List_toArrayAux___rarg(x_65, x_67); -x_69 = lean_compile_stage1(x_68, x_2, x_3, x_64); -if (lean_obj_tag(x_69) == 0) +x_92 = l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(x_1, x_2, x_3, x_86); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_72 = x_69; +lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_95 = x_92; } else { - lean_dec_ref(x_69); - x_72 = lean_box(0); + lean_dec_ref(x_92); + x_95 = lean_box(0); } -x_73 = lean_array_get_size(x_70); -x_74 = lean_usize_of_nat(x_73); -lean_dec(x_73); -x_75 = 0; -x_76 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; -x_77 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_76, x_70, x_74, x_75, x_76); -lean_dec(x_70); +x_96 = l_Lean_ConstantInfo_hasValue(x_93); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; +lean_dec(x_93); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -lean_dec(x_77); -if (lean_obj_tag(x_78) == 0) +x_97 = lean_box(0); +if (lean_is_scalar(x_95)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_95; +} +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_94); +return x_98; +} +else { -lean_object* x_79; lean_object* x_80; -x_79 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; -if (lean_is_scalar(x_72)) { - x_80 = lean_alloc_ctor(0, 2, 0); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_95); +x_99 = l_Lean_ConstantInfo_all(x_93); +lean_dec(x_93); +x_100 = l_List_redLength___rarg(x_99); +x_101 = lean_mk_empty_array_with_capacity(x_100); +lean_dec(x_100); +x_102 = l_List_toArrayAux___rarg(x_99, x_101); +x_103 = lean_compile_stage1(x_102, x_2, x_3, x_94); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; size_t x_108; size_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_106 = x_103; +} else { + lean_dec_ref(x_103); + x_106 = lean_box(0); +} +x_107 = lean_array_get_size(x_104); +x_108 = lean_usize_of_nat(x_107); +lean_dec(x_107); +x_109 = 0; +x_110 = l_Lean_Compiler_getStage1Decl_x3f___closed__1; +x_111 = l_Array_forInUnsafe_loop___at_Lean_Compiler_getStage1Decl_x3f___spec__4(x_1, x_110, x_104, x_108, x_109, x_110); +lean_dec(x_104); +lean_dec(x_1); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +lean_dec(x_111); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; +x_113 = l_Lean_Compiler_getStage1Decl_x3f___closed__3; +if (lean_is_scalar(x_106)) { + x_114 = lean_alloc_ctor(0, 2, 0); } else { - x_80 = x_72; + x_114 = x_106; } -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_71); -return x_80; +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_105); +return x_114; } else { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_78, 0); -lean_inc(x_81); -lean_dec(x_78); -if (lean_is_scalar(x_72)) { - x_82 = lean_alloc_ctor(0, 2, 0); +lean_object* x_115; lean_object* x_116; +x_115 = lean_ctor_get(x_112, 0); +lean_inc(x_115); +lean_dec(x_112); +if (lean_is_scalar(x_106)) { + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_82 = x_72; + x_116 = x_106; } -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_71); -return x_82; +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_105); +return x_116; } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_dec(x_1); -x_83 = lean_ctor_get(x_69, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_69, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_85 = x_69; +x_117 = lean_ctor_get(x_103, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_103, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_119 = x_103; } else { - lean_dec_ref(x_69); - x_85 = lean_box(0); + lean_dec_ref(x_103); + x_119 = lean_box(0); } -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); } else { - x_86 = x_85; + x_120 = x_119; +} +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_84); -return x_86; } } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_87 = lean_ctor_get(x_62, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_62, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_89 = x_62; +x_121 = lean_ctor_get(x_92, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_92, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_123 = x_92; } else { - lean_dec_ref(x_62); - x_89 = lean_box(0); + lean_dec_ref(x_92); + x_123 = lean_box(0); } -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); } else { - x_90 = x_89; + x_124 = x_123; } -lean_ctor_set(x_90, 0, x_87); -lean_ctor_set(x_90, 1, x_88); -return x_90; +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; } } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_91 = lean_ctor_get(x_61, 0); -lean_inc(x_91); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - x_92 = x_61; +x_125 = lean_ctor_get(x_91, 0); +lean_inc(x_125); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + x_126 = x_91; } else { - lean_dec_ref(x_61); - x_92 = lean_box(0); + lean_dec_ref(x_91); + x_126 = lean_box(0); } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 1, 0); } else { - x_93 = x_92; + x_127 = x_126; } -lean_ctor_set(x_93, 0, x_91); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_56); -return x_94; +lean_ctor_set(x_127, 0, x_125); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_86); +return x_128; } } } diff --git a/stage0/stdlib/Lean/Compiler/TerminalCases.c b/stage0/stdlib/Lean/Compiler/TerminalCases.c index 9830d368aa22..065f6b05502a 100644 --- a/stage0/stdlib/Lean/Compiler/TerminalCases.c +++ b/stage0/stdlib/Lean/Compiler/TerminalCases.c @@ -19,12 +19,11 @@ lean_object* l_Lean_Compiler_mkLetUsingScope(lean_object*, lean_object*, lean_ob LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___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_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___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_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2(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_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___closed__1; lean_object* l_Lean_Compiler_Decl_mapValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -lean_object* l_Lean_Compiler_withNewScopeImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -32,6 +31,7 @@ lean_object* l_Lean_mkAppN(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkLetDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); @@ -41,22 +41,20 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compile LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); lean_object* l_Lean_Compiler_attachJp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_mkLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1; +lean_object* l_Lean_Compiler_visitLambdaCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_visitLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg(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___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1; lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_Decl_terminalCases___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* l_Lean_Compiler_mkJpDeclIfNotSimple(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_mkLocalDecl(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_Decl_terminalCases(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_isCasesApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___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) { _start: @@ -265,7 +263,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -274,122 +272,452 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___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___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda(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, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; 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_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); -lean_dec(x_1); -x_8 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1; +lean_dec(x_6); +x_8 = lean_st_ref_get(x_2, 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_39 = lean_st_ref_get(x_4, x_10); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_st_ref_take(x_2, x_40); +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_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_45 = lean_ctor_get(x_42, 1); +lean_dec(x_45); +x_46 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1; +lean_ctor_set(x_42, 1, x_46); +x_47 = lean_st_ref_set(x_2, x_42, x_43); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Compiler_visitLambdaCore(x_1, x_2, x_3, x_4, x_48); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_50, 1); +lean_inc(x_53); +lean_dec(x_50); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_9 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(x_7, x_8, x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_9) == 0) +lean_inc(x_52); +x_54 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(x_53, x_52, x_2, x_3, x_4, x_51); +if (lean_obj_tag(x_54) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -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_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_Compiler_mkLetUsingScope(x_10, x_2, x_3, x_4, x_11); -if (lean_obj_tag(x_12) == 0) +x_57 = l_Lean_Compiler_mkLetUsingScope(x_55, x_2, x_3, x_4, x_56); +if (lean_obj_tag(x_57) == 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 = l_Lean_Compiler_mkLambda(x_6, x_13, x_2, x_3, x_4, x_14); -lean_dec(x_4); +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; uint8_t x_70; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = l_Lean_Compiler_mkLambda(x_52, x_58, x_2, x_3, x_4, x_59); lean_dec(x_3); +lean_dec(x_52); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_st_ref_get(x_4, x_62); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = lean_st_ref_get(x_2, x_64); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_ctor_get(x_9, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_9, 1); +lean_inc(x_69); +lean_dec(x_9); +x_70 = !lean_is_exclusive(x_66); +if (x_70 == 0) +{ +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_71 = lean_ctor_get(x_66, 1); +lean_dec(x_71); +x_72 = lean_ctor_get(x_66, 0); +lean_dec(x_72); +lean_ctor_set(x_66, 1, x_69); +lean_ctor_set(x_66, 0, x_68); +x_73 = lean_st_ref_get(x_4, x_67); +lean_dec(x_4); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = lean_st_ref_set(x_2, x_66, x_74); lean_dec(x_2); -lean_dec(x_6); -return x_15; +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) +{ +lean_object* x_77; +x_77 = lean_ctor_get(x_75, 0); +lean_dec(x_77); +lean_ctor_set(x_75, 0, x_61); +return x_75; } else { -uint8_t x_16; -lean_dec(x_6); +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_dec(x_75); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_61); +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; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_80 = lean_ctor_get(x_66, 2); +lean_inc(x_80); +lean_dec(x_66); +x_81 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_81, 0, x_68); +lean_ctor_set(x_81, 1, x_69); +lean_ctor_set(x_81, 2, x_80); +x_82 = lean_st_ref_get(x_4, x_67); lean_dec(x_4); -lean_dec(x_3); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +x_84 = lean_st_ref_set(x_2, x_81, x_83); lean_dec(x_2); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -return x_12; +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_86 = x_84; +} else { + lean_dec_ref(x_84); + x_86 = lean_box(0); +} +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_86; +} +lean_ctor_set(x_87, 0, x_61); +lean_ctor_set(x_87, 1, x_85); +return x_87; +} } 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_object* x_88; lean_object* x_89; +lean_dec(x_52); +lean_dec(x_3); +x_88 = lean_ctor_get(x_57, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_57, 1); +lean_inc(x_89); +lean_dec(x_57); +x_11 = x_88; +x_12 = x_89; +goto block_38; +} } +else +{ +lean_object* x_90; lean_object* x_91; +lean_dec(x_52); +lean_dec(x_3); +x_90 = lean_ctor_get(x_54, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_54, 1); +lean_inc(x_91); +lean_dec(x_54); +x_11 = x_90; +x_12 = x_91; +goto block_38; } } else { -uint8_t x_20; -lean_dec(x_6); -lean_dec(x_4); +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_92 = lean_ctor_get(x_42, 0); +x_93 = lean_ctor_get(x_42, 2); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_42); +x_94 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1; +x_95 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_st_ref_set(x_2, x_95, x_43); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Compiler_visitLambdaCore(x_1, x_2, x_3, x_4, x_97); +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_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_99, 1); +lean_inc(x_102); +lean_dec(x_99); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_101); +x_103 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(x_102, x_101, x_2, x_3, x_4, x_100); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_106 = l_Lean_Compiler_mkLetUsingScope(x_104, x_2, x_3, x_4, x_105); +if (lean_obj_tag(x_106) == 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; 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; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = l_Lean_Compiler_mkLambda(x_101, x_107, x_2, x_3, x_4, x_108); lean_dec(x_3); +lean_dec(x_101); +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 = lean_st_ref_get(x_4, x_111); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +lean_dec(x_112); +x_114 = lean_st_ref_get(x_2, x_113); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +x_117 = lean_ctor_get(x_9, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_9, 1); +lean_inc(x_118); +lean_dec(x_9); +x_119 = lean_ctor_get(x_115, 2); +lean_inc(x_119); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + lean_ctor_release(x_115, 2); + x_120 = x_115; +} else { + lean_dec_ref(x_115); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(0, 3, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_117); +lean_ctor_set(x_121, 1, x_118); +lean_ctor_set(x_121, 2, x_119); +x_122 = lean_st_ref_get(x_4, x_116); +lean_dec(x_4); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = lean_st_ref_set(x_2, x_121, x_123); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_9); -if (x_20 == 0) -{ -return x_9; +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_126)) { + x_127 = lean_alloc_ctor(0, 2, 0); +} else { + x_127 = x_126; +} +lean_ctor_set(x_127, 0, x_110); +lean_ctor_set(x_127, 1, x_125); +return x_127; } 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_object* x_128; lean_object* x_129; +lean_dec(x_101); +lean_dec(x_3); +x_128 = lean_ctor_get(x_106, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_106, 1); +lean_inc(x_129); +lean_dec(x_106); +x_11 = x_128; +x_12 = x_129; +goto block_38; } } +else +{ +lean_object* x_130; lean_object* x_131; +lean_dec(x_101); +lean_dec(x_3); +x_130 = lean_ctor_get(x_103, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_103, 1); +lean_inc(x_131); +lean_dec(x_103); +x_11 = x_130; +x_12 = x_131; +goto block_38; } } -static lean_object* _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1() { -_start: +block_38: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1), 5, 0); -return x_1; +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; +x_13 = lean_st_ref_get(x_4, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_get(x_2, x_14); +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_ctor_get(x_9, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_dec(x_9); +x_20 = !lean_is_exclusive(x_16); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_16, 1); +lean_dec(x_21); +x_22 = lean_ctor_get(x_16, 0); +lean_dec(x_22); +lean_ctor_set(x_16, 1, x_19); +lean_ctor_set(x_16, 0, x_18); +x_23 = lean_st_ref_get(x_4, x_17); +lean_dec(x_4); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_set(x_2, x_16, x_24); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +lean_ctor_set_tag(x_25, 1); +lean_ctor_set(x_25, 0, x_11); +return x_25; } +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_11); +lean_ctor_set(x_29, 1, x_28); +return x_29; } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +} +else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_alloc_closure((void*)(l_Lean_Compiler_visitLambda___boxed), 5, 1); -lean_closure_set(x_6, 0, x_1); -x_7 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1; -x_8 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_8, 0, x_6); -lean_closure_set(x_8, 1, x_7); -x_9 = l_Lean_Compiler_withNewScopeImp___rarg(x_8, x_2, x_3, x_4, x_5); -return x_9; +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_30 = lean_ctor_get(x_16, 2); +lean_inc(x_30); +lean_dec(x_16); +x_31 = lean_alloc_ctor(0, 3, 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); +x_32 = lean_st_ref_get(x_4, x_17); +lean_dec(x_4); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_st_ref_set(x_2, x_31, x_33); +lean_dec(x_2); +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); +} +if (lean_is_scalar(x_36)) { + x_37 = lean_alloc_ctor(1, 2, 0); +} else { + x_37 = x_36; + lean_ctor_set_tag(x_37, 1); +} +lean_ctor_set(x_37, 0, x_11); +lean_ctor_set(x_37, 1, x_35); +return x_37; +} +} } } LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___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) { @@ -407,7 +735,7 @@ x_16 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_vi return x_16; } } -static lean_object* _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -416,100 +744,6 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___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_inc(x_3); -x_8 = lean_array_push(x_1, x_3); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_9 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(x_2, x_8, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -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_6); -lean_inc(x_5); -lean_inc(x_4); -x_12 = l_Lean_Compiler_mkLetUsingScope(x_10, x_4, x_5, x_6, x_11); -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; -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_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1; -x_16 = lean_array_push(x_15, x_3); -x_17 = l_Lean_Compiler_mkLambda(x_16, x_13, x_4, x_5, x_6, x_14); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_16); -return x_17; -} -else -{ -uint8_t x_18; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) -{ -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -else -{ -uint8_t x_22; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(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: { @@ -605,150 +839,344 @@ return x_28; } else { -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_29; lean_object* x_30; lean_object* x_31; 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_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; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; x_29 = lean_ctor_get(x_14, 1); lean_inc(x_29); lean_dec(x_14); x_30 = lean_ctor_get(x_15, 0); lean_inc(x_30); lean_dec(x_15); -x_31 = 0; -x_32 = lean_box(x_31); -x_33 = lean_alloc_closure((void*)(l_Lean_Compiler_mkLocalDecl___boxed), 7, 3); -lean_closure_set(x_33, 0, x_7); -lean_closure_set(x_33, 1, x_12); -lean_closure_set(x_33, 2, x_32); -x_34 = lean_alloc_closure((void*)(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2), 7, 2); -lean_closure_set(x_34, 0, x_2); -lean_closure_set(x_34, 1, x_10); -x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Compiler_instMonadLCtxCompilerM___spec__2___rarg), 6, 2); -lean_closure_set(x_35, 0, x_33); -lean_closure_set(x_35, 1, x_34); +x_54 = lean_st_ref_get(x_5, x_29); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_get(x_3, 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); +lean_dec(x_56); +x_78 = lean_st_ref_get(x_5, x_58); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_st_ref_take(x_3, x_79); +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_ctor_get(x_81, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_81, 2); +lean_inc(x_84); +lean_dec(x_81); +x_85 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1; +x_86 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = lean_st_ref_set(x_3, x_86, x_82); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = 0; +x_90 = l_Lean_Compiler_mkLocalDecl(x_7, x_12, x_89, x_3, x_4, x_5, x_88); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +lean_inc(x_91); +x_93 = lean_array_push(x_2, x_91); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_36 = l_Lean_Compiler_withNewScopeImp___rarg(x_35, x_3, x_4, x_5, x_29); -if (lean_obj_tag(x_36) == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -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); +x_94 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet(x_10, x_93, x_3, x_4, x_5, x_92); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* 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); lean_inc(x_5); lean_inc(x_4); -x_39 = l_Lean_Compiler_mkJpDeclIfNotSimple(x_37, x_3, x_4, x_5, x_38); -if (lean_obj_tag(x_39) == 0) +lean_inc(x_3); +x_97 = l_Lean_Compiler_mkLetUsingScope(x_95, x_3, x_4, x_5, x_96); +if (lean_obj_tag(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; 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; uint8_t x_117; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1; +x_101 = lean_array_push(x_100, x_91); +x_102 = l_Lean_Compiler_mkLambda(x_101, x_98, x_3, x_4, x_5, x_99); +lean_dec(x_101); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = lean_st_ref_get(x_5, x_104); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_st_ref_get(x_3, x_106); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_ctor_get(x_57, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_57, 1); +lean_inc(x_111); +lean_dec(x_57); +x_112 = lean_ctor_get(x_108, 2); +lean_inc(x_112); +lean_dec(x_108); +x_113 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_113, 0, x_110); +lean_ctor_set(x_113, 1, x_111); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_st_ref_get(x_5, x_109); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_st_ref_set(x_3, x_113, x_115); +x_117 = !lean_is_exclusive(x_116); +if (x_117 == 0) +{ +lean_object* x_118; +x_118 = lean_ctor_get(x_116, 0); +lean_dec(x_118); +lean_ctor_set(x_116, 0, x_103); +x_31 = x_116; +goto block_53; +} +else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -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); +lean_object* x_119; lean_object* x_120; +x_119 = lean_ctor_get(x_116, 1); +lean_inc(x_119); +lean_dec(x_116); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_103); +lean_ctor_set(x_120, 1, x_119); +x_31 = x_120; +goto block_53; +} +} +else +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_91); +x_121 = lean_ctor_get(x_97, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_97, 1); +lean_inc(x_122); +lean_dec(x_97); +x_59 = x_121; +x_60 = x_122; +goto block_77; +} +} +else +{ +lean_object* x_123; lean_object* x_124; +lean_dec(x_91); +x_123 = lean_ctor_get(x_94, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_94, 1); +lean_inc(x_124); +lean_dec(x_94); +x_59 = x_123; +x_60 = x_124; +goto block_77; +} +block_53: +{ +if (lean_obj_tag(x_31) == 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); +lean_inc(x_5); +lean_inc(x_4); +x_34 = l_Lean_Compiler_mkJpDeclIfNotSimple(x_32, x_3, x_4, x_5, x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +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); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_42 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases(x_30, x_13, x_3, x_4, x_5, x_41); -if (lean_obj_tag(x_42) == 0) +x_37 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases(x_30, x_13, x_3, x_4, x_5, x_36); +if (lean_obj_tag(x_37) == 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 = l_Lean_Compiler_attachJp(x_43, x_40, x_3, x_4, x_5, x_44); -return x_45; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = l_Lean_Compiler_attachJp(x_38, x_35, x_3, x_4, x_5, x_39); +return x_40; } else { -uint8_t x_46; -lean_dec(x_40); +uint8_t x_41; +lean_dec(x_35); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_46 = !lean_is_exclusive(x_42); -if (x_46 == 0) +x_41 = !lean_is_exclusive(x_37); +if (x_41 == 0) { -return x_42; +return x_37; } 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_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_37, 0); +x_43 = lean_ctor_get(x_37, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_37); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } else { -uint8_t x_50; +uint8_t x_45; lean_dec(x_30); lean_dec(x_13); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_50 = !lean_is_exclusive(x_39); -if (x_50 == 0) +x_45 = !lean_is_exclusive(x_34); +if (x_45 == 0) { -return x_39; +return x_34; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_39, 0); -x_52 = lean_ctor_get(x_39, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_39); -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; +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_34, 0); +x_47 = lean_ctor_get(x_34, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_34); +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; } } } else { -uint8_t x_54; +uint8_t x_49; lean_dec(x_30); lean_dec(x_13); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_54 = !lean_is_exclusive(x_36); -if (x_54 == 0) +x_49 = !lean_is_exclusive(x_31); +if (x_49 == 0) { -return x_36; +return x_31; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_36, 0); -x_56 = lean_ctor_get(x_36, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_36); -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; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_31, 0); +x_51 = lean_ctor_get(x_31, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_31); +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_77: +{ +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; uint8_t x_73; +x_61 = lean_st_ref_get(x_5, x_60); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec(x_61); +x_63 = lean_st_ref_get(x_3, x_62); +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_57, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_57, 1); +lean_inc(x_67); +lean_dec(x_57); +x_68 = lean_ctor_get(x_64, 2); +lean_inc(x_68); +lean_dec(x_64); +x_69 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_69, 0, x_66); +lean_ctor_set(x_69, 1, x_67); +lean_ctor_set(x_69, 2, x_68); +x_70 = lean_st_ref_get(x_5, x_65); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = lean_st_ref_set(x_3, x_69, x_71); +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) +{ +lean_object* x_74; +x_74 = lean_ctor_get(x_72, 0); +lean_dec(x_74); +lean_ctor_set_tag(x_72, 1); +lean_ctor_set(x_72, 0, x_59); +x_31 = x_72; +goto block_53; +} +else +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +lean_dec(x_72); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_59); +lean_ctor_set(x_76, 1, x_75); +x_31 = x_76; +goto block_53; } } } } else { -uint8_t x_58; +uint8_t x_125; lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); @@ -757,105 +1185,105 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_58 = !lean_is_exclusive(x_14); -if (x_58 == 0) +x_125 = !lean_is_exclusive(x_14); +if (x_125 == 0) { return x_14; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_14, 0); -x_60 = lean_ctor_get(x_14, 1); -lean_inc(x_60); -lean_inc(x_59); +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_14, 0); +x_127 = lean_ctor_get(x_14, 1); +lean_inc(x_127); +lean_inc(x_126); lean_dec(x_14); -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; +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; } } } else { -lean_object* x_62; lean_object* x_63; -x_62 = lean_expr_instantiate_rev(x_1, x_2); +lean_object* x_129; lean_object* x_130; +x_129 = lean_expr_instantiate_rev(x_1, x_2); lean_dec(x_2); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_62); -x_63 = l_Lean_Compiler_isCasesApp_x3f(x_62, x_4, x_5, x_6); -if (lean_obj_tag(x_63) == 0) +lean_inc(x_129); +x_130 = l_Lean_Compiler_isCasesApp_x3f(x_129, x_4, x_5, x_6); +if (lean_obj_tag(x_130) == 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_131; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +if (lean_obj_tag(x_131) == 0) { -uint8_t x_65; +uint8_t x_132; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_65 = !lean_is_exclusive(x_63); -if (x_65 == 0) +x_132 = !lean_is_exclusive(x_130); +if (x_132 == 0) { -lean_object* x_66; -x_66 = lean_ctor_get(x_63, 0); -lean_dec(x_66); -lean_ctor_set(x_63, 0, x_62); -return x_63; +lean_object* x_133; +x_133 = lean_ctor_get(x_130, 0); +lean_dec(x_133); +lean_ctor_set(x_130, 0, x_129); +return x_130; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_63, 1); -lean_inc(x_67); -lean_dec(x_63); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_62); -lean_ctor_set(x_68, 1, x_67); -return x_68; +lean_object* x_134; lean_object* x_135; +x_134 = lean_ctor_get(x_130, 1); +lean_inc(x_134); +lean_dec(x_130); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_129); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_63, 1); -lean_inc(x_69); -lean_dec(x_63); -x_70 = lean_ctor_get(x_64, 0); -lean_inc(x_70); -lean_dec(x_64); -x_71 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases(x_70, x_62, x_3, x_4, x_5, x_69); -return x_71; +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_130, 1); +lean_inc(x_136); +lean_dec(x_130); +x_137 = lean_ctor_get(x_131, 0); +lean_inc(x_137); +lean_dec(x_131); +x_138 = l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases(x_137, x_129, x_3, x_4, x_5, x_136); +return x_138; } } else { -uint8_t x_72; -lean_dec(x_62); +uint8_t x_139; +lean_dec(x_129); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_72 = !lean_is_exclusive(x_63); -if (x_72 == 0) +x_139 = !lean_is_exclusive(x_130); +if (x_139 == 0) { -return x_63; +return x_130; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_63, 0); -x_74 = lean_ctor_get(x_63, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_63); -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; +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_130, 0); +x_141 = lean_ctor_get(x_130, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_130); +x_142 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +return x_142; } } } @@ -926,12 +1354,10 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___closed__1 = _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitCases___closed__1); -l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1 = _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___lambda__1___closed__1); l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1 = _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1(); lean_mark_persistent(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLambda___closed__1); -l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1 = _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___lambda__2___closed__1); +l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1 = _init_l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1(); +lean_mark_persistent(l___private_Lean_Compiler_TerminalCases_0__Lean_Compiler_TerminalCases_visitLet___closed__1); l_Lean_Compiler_Decl_terminalCases___closed__1 = _init_l_Lean_Compiler_Decl_terminalCases___closed__1(); lean_mark_persistent(l_Lean_Compiler_Decl_terminalCases___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Compiler/Util.c b/stage0/stdlib/Lean/Compiler/Util.c index b6ed11520f97..47798dc83766 100644 --- a/stage0/stdlib/Lean/Compiler/Util.c +++ b/stage0/stdlib/Lean/Compiler/Util.c @@ -14,29 +14,41 @@ extern "C" { #endif static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_isRuntimeBultinType(lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__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_LocalDecl_userName(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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_Compiler_getLambdaArity___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_isLCProof(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_inc(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_InductiveVal_numCtors(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Lean_Compiler_lcnfSizeLe_go___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__21; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__38; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__2; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__3; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isLcCast_x3f(lean_object*); static lean_object* l_Lean_Compiler_isCasesApp_x3f___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__42; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__1; @@ -44,18 +56,25 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_g static lean_object* l_Lean_Compiler_isLcUnreachable___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_isLcUnreachable___boxed(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__33; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Compiler_Util_0__Lean_Compiler_getCasesOnInductiveVal_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_isLcCast_x3f___closed__2; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__18; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_getCasesInfo_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_isLcUnreachable___closed__1; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_inc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_isLcUnreachable(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__22; LEAN_EXPORT lean_object* l_Lean_Compiler_getCtorArity_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__14; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__5; static lean_object* l_Lean_Compiler_isCasesApp_x3f___closed__4; @@ -63,6 +82,7 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__30; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__28; lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isJpBinderName___boxed(lean_object*); lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_isJpBinderName___closed__1; @@ -76,36 +96,50 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_g lean_object* l_Lean_Core_instInhabitedCoreM___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__8; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_object*); static lean_object* l_Lean_Compiler_mkLcProof___closed__1; extern lean_object* l_Lean_levelZero; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Lean_Compiler_lcnfSizeLe_go___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CasesInfo_geNumDiscrs___boxed(lean_object*); static lean_object* l_Lean_Compiler_isLcCast_x3f___closed__1; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__7; +static lean_object* l_Lean_Compiler_isEmptyType___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_CasesInfo_updateResultingType___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isLcCast_x3f___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__29; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__24; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__15; lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_isEmptyType___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CasesInfo_updateResultingType_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_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_st_ref_take(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_isEmptyType___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_findDecl_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__37; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_getCasesInfo_x3f___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getCasesInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__20; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_Compiler_builtinRuntimeTypes___closed__32; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__10; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLCNFSize_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_getCasesInfo_x3f___spec__2___closed__4; lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); static lean_object* l_Lean_Compiler_getLCNFSize_go___closed__1; @@ -117,28 +151,41 @@ lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getCtorArity_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_mkLcProof(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLambdaArity(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isRuntimeBultinType___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__1___boxed(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_EXPORT lean_object* l_Lean_Compiler_CasesInfo_geNumDiscrs(lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_getCasesInfo_x3f___spec__2___closed__1; uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__19; LEAN_EXPORT lean_object* l_Lean_Compiler_isCompilerRelevantMData___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__39; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_getCasesInfo_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_onlyOneExitPoint(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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*); uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isSimpleLCNF(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__31; LEAN_EXPORT lean_object* l_Lean_LocalDecl_isJp___boxed(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__11; LEAN_EXPORT lean_object* l_Lean_Compiler_isLCProof___boxed(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_isJpBinderName(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__4(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_Compiler_isCompilerRelevantMData(lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__12; @@ -155,16 +202,23 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Compiler_isCasesApp_x3f___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__40; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__9; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__9(lean_object*, 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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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*, lean_object*); static lean_object* l_Lean_Compiler_isLCProof___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_getLCNFSize_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_manyExitPoints(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__13; +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__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_Compiler_builtinRuntimeTypes___closed__23; static lean_object* l_Lean_Compiler_isLCProof___closed__1; @@ -183,14 +237,17 @@ lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getLCNFSize(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CasesInfo_updateResultingType_go___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___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_Compiler_lcnfSizeLe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__4; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__27; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__7___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_loop___at_Lean_Compiler_getLCNFSize_go___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_Lean_Compiler_isEmptyType(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_isJump_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_getCasesInfo_x3f___spec__2___closed__3; static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__26; uint8_t l_Lean_Expr_isLet(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_builtinRuntimeTypes___closed__1; LEAN_EXPORT uint8_t l_Lean_Compiler_isCompilerRelevantMData(lean_object* x_1) { @@ -5599,6 +5656,5334 @@ return x_23; } } } +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_inc(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; uint8_t x_18; +x_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_st_ref_take(x_2, 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_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_9, x_11); +lean_dec(x_9); +x_13 = lean_st_ref_set(x_2, x_12, x_10); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_get(x_4, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_2, x_16); +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 = lean_nat_dec_le(x_19, x_1); +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 = lean_nat_dec_le(x_22, x_1); +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_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_inc___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_Compiler_lcnfSizeLe_inc(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__1() { +_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_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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_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_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +lean_dec(x_4); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_5, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4; +x_21 = l_panic___at_Lean_Expr_getRevArg_x21___spec__1(x_20); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_22 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_21, x_9, x_10, x_11, x_12); +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_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +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; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_32; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_31; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_array_fget(x_2, x_5); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_39 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_38, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +lean_ctor_set(x_39, 0, x_44); +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +x_46 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +lean_inc(x_3); +{ +lean_object* _tmp_3 = x_17; +lean_object* _tmp_4 = x_49; +lean_object* _tmp_7 = x_3; +lean_object* _tmp_11 = x_48; +x_4 = _tmp_3; +x_5 = _tmp_4; +x_8 = _tmp_7; +x_12 = _tmp_11; +} +goto _start; +} +} +else +{ +uint8_t x_51; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_39); +if (x_51 == 0) +{ +return x_39; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_39, 0); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_39); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +lean_object* x_55; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_8); +lean_ctor_set(x_55, 1, x_12); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_8); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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_object* x_7; lean_object* x_8; +x_6 = 1; +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_5); +return x_8; +} +} +static lean_object* _init_l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Compiler_lcnfSizeLe_go___lambda__1___boxed), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_lcnfSizeLe_go___lambda__2___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(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_Compiler_lcnfSizeLe_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_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__2(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__3(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__4(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__5(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__6(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__7(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__8(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___lambda__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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__9(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_go___lambda__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) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_8 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1; +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_5(x_11, x_12, x_4, x_5, x_6, x_10); +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_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_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l_Lean_Compiler_getLCNFSize_go___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_18, x_20); +x_22 = lean_ctor_get(x_14, 3); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 2); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +x_27 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__10(x_2, x_21, x_26, x_23, x_24, x_23, x_25, x_26, x_4, x_5, x_6, x_10); +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +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 = lean_apply_5(x_11, x_31, x_4, x_5, x_6, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_27, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_35); +return x_27; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_27); +if (x_39 == 0) +{ +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 0); +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_27); +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 +{ +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_8); +if (x_43 == 0) +{ +return x_8; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_8, 0); +x_45 = lean_ctor_get(x_8, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_8); +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_Lean_Compiler_lcnfSizeLe_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: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = 0; +x_13 = lean_box(x_12); +lean_ctor_set(x_7, 0, x_13); +return x_7; +} +else +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_7, 1); +lean_inc(x_14); +lean_dec(x_7); +x_15 = 0; +x_16 = lean_box(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_14); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_7, 1); +lean_inc(x_18); +lean_dec(x_7); +x_19 = lean_box(0); +x_20 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2(x_2, x_1, x_19, x_3, x_4, x_5, x_18); +return x_20; +} +} +case 1: +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +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) +{ +uint8_t x_24; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = 0; +x_27 = lean_box(x_26); +lean_ctor_set(x_21, 0, x_27); +return x_21; +} +else +{ +lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_21, 1); +lean_inc(x_28); +lean_dec(x_21); +x_29 = 0; +x_30 = lean_box(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_28); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_21, 1); +lean_inc(x_32); +lean_dec(x_21); +x_33 = lean_box(0); +x_34 = l_Lean_Compiler_lcnfSizeLe_go___lambda__3(x_2, x_1, x_33, x_3, x_4, x_5, x_32); +return x_34; +} +} +case 2: +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) +{ +uint8_t x_38; +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_35); +if (x_38 == 0) +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 0); +lean_dec(x_39); +x_40 = 0; +x_41 = lean_box(x_40); +lean_ctor_set(x_35, 0, x_41); +return x_35; +} +else +{ +lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_35, 1); +lean_inc(x_42); +lean_dec(x_35); +x_43 = 0; +x_44 = lean_box(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_42); +return x_45; +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_35, 1); +lean_inc(x_46); +lean_dec(x_35); +x_47 = lean_box(0); +x_48 = l_Lean_Compiler_lcnfSizeLe_go___lambda__4(x_2, x_1, x_47, x_3, x_4, x_5, x_46); +return x_48; +} +} +case 3: +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_unbox(x_50); +lean_dec(x_50); +if (x_51 == 0) +{ +uint8_t x_52; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_52 = !lean_is_exclusive(x_49); +if (x_52 == 0) +{ +lean_object* x_53; uint8_t x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_49, 0); +lean_dec(x_53); +x_54 = 0; +x_55 = lean_box(x_54); +lean_ctor_set(x_49, 0, x_55); +return x_49; +} +else +{ +lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_49, 1); +lean_inc(x_56); +lean_dec(x_49); +x_57 = 0; +x_58 = lean_box(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_56); +return x_59; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_49, 1); +lean_inc(x_60); +lean_dec(x_49); +x_61 = lean_box(0); +x_62 = l_Lean_Compiler_lcnfSizeLe_go___lambda__5(x_2, x_1, x_61, x_3, x_4, x_5, x_60); +return x_62; +} +} +case 4: +{ +lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_63 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_unbox(x_64); +lean_dec(x_64); +if (x_65 == 0) +{ +uint8_t x_66; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_63); +if (x_66 == 0) +{ +lean_object* x_67; uint8_t x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_63, 0); +lean_dec(x_67); +x_68 = 0; +x_69 = lean_box(x_68); +lean_ctor_set(x_63, 0, x_69); +return x_63; +} +else +{ +lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_63, 1); +lean_inc(x_70); +lean_dec(x_63); +x_71 = 0; +x_72 = lean_box(x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_70); +return x_73; +} +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_63, 1); +lean_inc(x_74); +lean_dec(x_63); +x_75 = lean_box(0); +x_76 = l_Lean_Compiler_lcnfSizeLe_go___lambda__6(x_2, x_1, x_75, x_3, x_4, x_5, x_74); +return x_76; +} +} +case 5: +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_unbox(x_78); +lean_dec(x_78); +if (x_79 == 0) +{ +uint8_t x_80; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_77); +if (x_80 == 0) +{ +lean_object* x_81; uint8_t x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_77, 0); +lean_dec(x_81); +x_82 = 0; +x_83 = lean_box(x_82); +lean_ctor_set(x_77, 0, x_83); +return x_77; +} +else +{ +lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_77, 1); +lean_inc(x_84); +lean_dec(x_77); +x_85 = 0; +x_86 = lean_box(x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_84); +return x_87; +} +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_77, 1); +lean_inc(x_88); +lean_dec(x_77); +x_89 = lean_box(0); +x_90 = l_Lean_Compiler_lcnfSizeLe_go___lambda__7(x_2, x_1, x_89, x_3, x_4, x_5, x_88); +return x_90; +} +} +case 6: +{ +lean_object* x_91; +x_91 = lean_ctor_get(x_2, 2); +lean_inc(x_91); +lean_dec(x_2); +x_2 = x_91; +goto _start; +} +case 7: +{ +lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_93 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_unbox(x_94); +lean_dec(x_94); +if (x_95 == 0) +{ +uint8_t x_96; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_96 = !lean_is_exclusive(x_93); +if (x_96 == 0) +{ +lean_object* x_97; uint8_t x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_93, 0); +lean_dec(x_97); +x_98 = 0; +x_99 = lean_box(x_98); +lean_ctor_set(x_93, 0, x_99); +return x_93; +} +else +{ +lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; +x_100 = lean_ctor_get(x_93, 1); +lean_inc(x_100); +lean_dec(x_93); +x_101 = 0; +x_102 = lean_box(x_101); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +return x_103; +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_93, 1); +lean_inc(x_104); +lean_dec(x_93); +x_105 = lean_box(0); +x_106 = l_Lean_Compiler_lcnfSizeLe_go___lambda__8(x_2, x_1, x_105, x_3, x_4, x_5, x_104); +return x_106; +} +} +case 8: +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; +x_107 = lean_ctor_get(x_2, 2); +lean_inc(x_107); +x_108 = lean_ctor_get(x_2, 3); +lean_inc(x_108); +lean_dec(x_2); +x_109 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_unbox(x_110); +if (x_111 == 0) +{ +uint8_t x_112; +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_112 = !lean_is_exclusive(x_109); +if (x_112 == 0) +{ +lean_object* x_113; +x_113 = lean_ctor_get(x_109, 0); +lean_dec(x_113); +return x_109; +} +else +{ +lean_object* x_114; lean_object* x_115; +x_114 = lean_ctor_get(x_109, 1); +lean_inc(x_114); +lean_dec(x_109); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_110); +lean_ctor_set(x_115, 1, x_114); +return x_115; +} +} +else +{ +lean_object* x_116; lean_object* x_117; +lean_dec(x_110); +x_116 = lean_ctor_get(x_109, 1); +lean_inc(x_116); +lean_dec(x_109); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_117 = l_Lean_Compiler_lcnfSizeLe_go(x_1, x_107, x_3, x_4, x_5, x_116); +if (lean_obj_tag(x_117) == 0) +{ +lean_object* x_118; uint8_t x_119; +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_unbox(x_118); +if (x_119 == 0) +{ +uint8_t x_120; +lean_dec(x_108); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_120 = !lean_is_exclusive(x_117); +if (x_120 == 0) +{ +lean_object* x_121; +x_121 = lean_ctor_get(x_117, 0); +lean_dec(x_121); +return x_117; +} +else +{ +lean_object* x_122; lean_object* x_123; +x_122 = lean_ctor_get(x_117, 1); +lean_inc(x_122); +lean_dec(x_117); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_118); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +else +{ +lean_object* x_124; +lean_dec(x_118); +x_124 = lean_ctor_get(x_117, 1); +lean_inc(x_124); +lean_dec(x_117); +x_2 = x_108; +x_6 = x_124; +goto _start; +} +} +else +{ +uint8_t x_126; +lean_dec(x_108); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_126 = !lean_is_exclusive(x_117); +if (x_126 == 0) +{ +return x_117; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_117, 0); +x_128 = lean_ctor_get(x_117, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_117); +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; +} +} +} +} +case 9: +{ +lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_130 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +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) +{ +uint8_t x_133; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_133 = !lean_is_exclusive(x_130); +if (x_133 == 0) +{ +lean_object* x_134; uint8_t x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_130, 0); +lean_dec(x_134); +x_135 = 0; +x_136 = lean_box(x_135); +lean_ctor_set(x_130, 0, x_136); +return x_130; +} +else +{ +lean_object* x_137; uint8_t x_138; lean_object* x_139; lean_object* x_140; +x_137 = lean_ctor_get(x_130, 1); +lean_inc(x_137); +lean_dec(x_130); +x_138 = 0; +x_139 = lean_box(x_138); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_137); +return x_140; +} +} +else +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_130, 1); +lean_inc(x_141); +lean_dec(x_130); +x_142 = lean_box(0); +x_143 = l_Lean_Compiler_lcnfSizeLe_go___lambda__9(x_2, x_1, x_142, x_3, x_4, x_5, x_141); +return x_143; +} +} +case 10: +{ +lean_object* x_144; lean_object* x_145; uint8_t x_146; +x_144 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_unbox(x_145); +lean_dec(x_145); +if (x_146 == 0) +{ +uint8_t x_147; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_147 = !lean_is_exclusive(x_144); +if (x_147 == 0) +{ +lean_object* x_148; uint8_t x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_144, 0); +lean_dec(x_148); +x_149 = 0; +x_150 = lean_box(x_149); +lean_ctor_set(x_144, 0, x_150); +return x_144; +} +else +{ +lean_object* x_151; uint8_t x_152; lean_object* x_153; lean_object* x_154; +x_151 = lean_ctor_get(x_144, 1); +lean_inc(x_151); +lean_dec(x_144); +x_152 = 0; +x_153 = lean_box(x_152); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +return x_154; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_144, 1); +lean_inc(x_155); +lean_dec(x_144); +x_156 = lean_box(0); +x_157 = l_Lean_Compiler_lcnfSizeLe_go___lambda__10(x_2, x_1, x_156, x_3, x_4, x_5, x_155); +return x_157; +} +} +default: +{ +lean_object* x_158; lean_object* x_159; uint8_t x_160; +x_158 = l_Lean_Compiler_lcnfSizeLe_inc(x_1, x_3, x_4, x_5, x_6); +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_unbox(x_159); +lean_dec(x_159); +if (x_160 == 0) +{ +uint8_t x_161; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_161 = !lean_is_exclusive(x_158); +if (x_161 == 0) +{ +lean_object* x_162; uint8_t x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_158, 0); +lean_dec(x_162); +x_163 = 0; +x_164 = lean_box(x_163); +lean_ctor_set(x_158, 0, x_164); +return x_158; +} +else +{ +lean_object* x_165; uint8_t x_166; lean_object* x_167; lean_object* x_168; +x_165 = lean_ctor_get(x_158, 1); +lean_inc(x_165); +lean_dec(x_158); +x_166 = 0; +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_165); +return x_168; +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_158, 1); +lean_inc(x_169); +lean_dec(x_158); +x_170 = lean_box(0); +x_171 = l_Lean_Compiler_lcnfSizeLe_go___lambda__11(x_2, x_1, x_170, x_3, x_4, x_5, x_169); +return x_171; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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, 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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_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, 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_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__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); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__6(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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__7(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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__8(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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__9(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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__10(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_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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_Compiler_lcnfSizeLe_go___lambda__1(x_1, x_2, x_3, x_4, 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_Lean_Compiler_lcnfSizeLe_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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe_go___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_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Compiler_lcnfSizeLe_go___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_lcnfSizeLe(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_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_st_mk_ref(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); +lean_dec(x_9); +lean_inc(x_4); +lean_inc(x_10); +x_12 = l_Lean_Compiler_lcnfSizeLe_go(x_2, x_1, x_10, x_3, x_4, x_11); +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; +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_4, x_14); +lean_dec(x_4); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_10, x_16); +lean_dec(x_10); +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_13); +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_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_10); +lean_dec(x_4); +x_22 = !lean_is_exclusive(x_12); +if (x_22 == 0) +{ +return x_12; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_12); +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_Lean_Compiler_getLambdaArity(lean_object* x_1) { _start: { @@ -5837,6 +11222,275 @@ lean_dec(x_3); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_manyExitPoints(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 6: +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +lean_dec(x_1); +x_1 = x_5; +goto _start; +} +case 8: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 3); +lean_inc(x_8); +lean_dec(x_1); +x_9 = l_Lean_Compiler_isJpBinderName(x_7); +lean_dec(x_7); +if (x_9 == 0) +{ +x_1 = x_8; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +x_11 = lean_box(x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_4); +return x_12; +} +} +default: +{ +lean_object* x_13; +x_13 = l_Lean_Compiler_isCasesApp_x3f(x_1, x_2, x_3, x_4); +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; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 0); +lean_dec(x_16); +x_17 = 0; +x_18 = lean_box(x_17); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; uint8_t 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 = 0; +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_14); +x_23 = !lean_is_exclusive(x_13); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_13, 0); +lean_dec(x_24); +x_25 = 1; +x_26 = lean_box(x_25); +lean_ctor_set(x_13, 0, x_26); +return x_13; +} +else +{ +lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_13, 1); +lean_inc(x_27); +lean_dec(x_13); +x_28 = 1; +x_29 = lean_box(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_27); +return x_30; +} +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_13); +if (x_31 == 0) +{ +return x_13; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_13, 0); +x_33 = lean_ctor_get(x_13, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_13); +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_Lean_Compiler_onlyOneExitPoint(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_Compiler_manyExitPoints(x_1, x_2, x_3, x_4); +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); +lean_dec(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_5, 0); +lean_dec(x_9); +x_10 = 1; +x_11 = lean_box(x_10); +lean_ctor_set(x_5, 0, x_11); +return x_5; +} +else +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_5, 1); +lean_inc(x_12); +lean_dec(x_5); +x_13 = 1; +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); +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_5); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_5, 0); +lean_dec(x_17); +x_18 = 0; +x_19 = lean_box(x_18); +lean_ctor_set(x_5, 0, x_19); +return x_5; +} +else +{ +lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_dec(x_5); +x_21 = 0; +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_20); +return x_23; +} +} +} +else +{ +uint8_t x_24; +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; +} +} +} +} +static lean_object* _init_l_Lean_Compiler_isEmptyType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Empty", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_isEmptyType___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Compiler_isEmptyType___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_isEmptyType(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_object* x_7; lean_object* x_8; +x_5 = l_Lean_Compiler_isEmptyType___closed__2; +x_6 = l_Lean_Expr_isConstOf(x_1, x_5); +x_7 = lean_box(x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_isEmptyType___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_Compiler_isEmptyType(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} lean_object* initialize_Init(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Match_MatcherInfo(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -5984,6 +11638,18 @@ l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4 = lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_getLCNFSize_go___spec__1___closed__4); l_Lean_Compiler_getLCNFSize_go___closed__1 = _init_l_Lean_Compiler_getLCNFSize_go___closed__1(); lean_mark_persistent(l_Lean_Compiler_getLCNFSize_go___closed__1); +l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__1); +l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Compiler_lcnfSizeLe_go___spec__1___closed__2); +l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1 = _init_l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__1); +l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2 = _init_l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Compiler_lcnfSizeLe_go___lambda__2___closed__2); +l_Lean_Compiler_isEmptyType___closed__1 = _init_l_Lean_Compiler_isEmptyType___closed__1(); +lean_mark_persistent(l_Lean_Compiler_isEmptyType___closed__1); +l_Lean_Compiler_isEmptyType___closed__2 = _init_l_Lean_Compiler_isEmptyType___closed__2(); +lean_mark_persistent(l_Lean_Compiler_isEmptyType___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/Quotation.c b/stage0/stdlib/Lean/Elab/Quotation.c index 8c111b8a97f7..bc497c3327a2 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -15,26 +15,25 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); uint8_t l_Lean_Syntax_isQuot(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18___boxed(lean_object**); 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*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__50; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__18; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__3; lean_object* l_Lean_Elab_Term_Quotation_getAntiquotKindSpec_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2; lean_object* l_Lean_Syntax_unescapeAntiquot(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9___lambda__1___closed__1; @@ -43,6 +42,7 @@ uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(lean_object*, lean_object*, 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_Quotation_stxQuot_expand___closed__38; lean_object* l_Lean_extractMacroScopes(lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__17; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6; size_t lean_usize_add(size_t, size_t); @@ -50,13 +50,13 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__21; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__29; lean_object* l_List_tail_x21___rarg(lean_object*); uint8_t l_Lean_Syntax_isTokenAntiquot(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__6(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__48; @@ -75,26 +75,27 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__4; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___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*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__30; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__21; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4; lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__15; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6; @@ -111,14 +112,17 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__16; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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* l_Array_append___rarg(lean_object*, lean_object*); @@ -127,7 +131,6 @@ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation_____ lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(lean_object*, lean_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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable___boxed(lean_object*, lean_object*); @@ -136,11 +139,12 @@ static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__34; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(lean_object*, 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*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(lean_object*, lean_object*); lean_object* l_Lean_Syntax_antiquotSuffixSplice_x3f(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; lean_object* l_Lean_SourceInfo_fromRef(lean_object*); @@ -148,45 +152,46 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__6___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4; lean_object* l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__1; lean_object* l_List_unzip___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__16___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_828____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__7(size_t, size_t, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_append(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___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_Term_Quotation_getQuotKind___closed__16; size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___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_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__40; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__9; @@ -194,7 +199,6 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__6; @@ -207,57 +211,47 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambd static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__5(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_Quotation_stxQuot_expand___closed__27; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__10(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__24; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___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_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19___boxed(lean_object*, 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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__12; -static lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__11; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__47; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1; lean_object* lean_array_get_size(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +static lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__8(size_t, 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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__23; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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*, lean_object*, lean_object*); lean_object* l_List_range(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20; lean_object* l_Lean_MessageData_ofList(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__7___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*); @@ -266,30 +260,32 @@ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation_____ lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1; lean_object* l_Lean_Syntax_getAntiquotTerm(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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*, 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_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; lean_object* l_Std_mkHashSetImp___rarg(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___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_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_Quotation_getQuotKind___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_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__44; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__6; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; @@ -303,6 +299,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_panic___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__1___closed__1; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange(lean_object*); lean_object* l_Lean_Elab_Term_Quotation_getPatternsVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__7; @@ -316,22 +314,26 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__13; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__3; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__41; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__9; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__23; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand(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_Quotation_0__Lean_Elab_Term_Quotation_findUsedAlts_go___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; @@ -339,61 +341,57 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___boxed(lean_object**); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_(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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__21; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__3; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__18; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20___boxed(lean_object**); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_Quotation_getQuotKind___closed__23; LEAN_EXPORT lean_object* l_Std_mkHashSet___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms(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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__30; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11; @@ -403,12 +401,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16; extern lean_object* l_Lean_LocalContext_empty; size_t lean_uint64_to_usize(uint64_t); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(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_Quotation_stxQuot_expand___closed__42; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__1; @@ -416,37 +412,31 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambd LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___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_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__18; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4; uint8_t l_Lean_Syntax_isAntiquots(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_EXPORT lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot__; lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__15; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4; LEAN_EXPORT lean_object* l_List_replace___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__4; lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getQuotKind___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25(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*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_docString___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__1; @@ -456,10 +446,10 @@ lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18(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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__16(lean_object*, lean_object*, 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2; @@ -467,13 +457,14 @@ static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Ela static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__37; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(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_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__18; @@ -489,6 +480,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__25; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__9; @@ -502,7 +495,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotati LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__45; -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; @@ -512,12 +504,10 @@ static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Ela static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__3; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Quotation_getQuotKind___spec__1___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_Quotation_elabNoErrorIfUnused_declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__4; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__17; @@ -532,6 +522,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__17; +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_docString(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; @@ -545,6 +536,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__8; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; @@ -554,7 +546,6 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ lean_object* l_Lean_Elab_Term_elabTerm(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_Quotation_getQuotKind___closed__25; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4; LEAN_EXPORT uint8_t l_Std_HashSetImp_contains___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts___spec__1(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; @@ -571,47 +562,50 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* l_Lean_Unhygienic_run___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__3; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__13; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_findUsedAlts_go(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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__9; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9(lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__2; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__22; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___closed__1; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Std_HashSetImp_moveEntries___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Quotation_getQuotKind___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__3; lean_object* l_Std_RBNode_find___at_Lean_sanitizeName___spec__1(lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__2; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__3; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__3; lean_object* l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__7; @@ -620,7 +614,6 @@ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation_____ static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__2; lean_object* l_Nat_repr(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__15; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_Quotation_getQuotKind___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -634,25 +627,29 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22___boxed(lean_object**); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9(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_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__5; 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__6; uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_828____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_367____spec__1(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__6; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__21; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__52; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6(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_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__5; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__28; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__6; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__8; @@ -666,8 +663,6 @@ lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_Quotation_getQuotKind___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts___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_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; @@ -677,12 +672,12 @@ lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__13(lean_object*, lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3; lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1; uint8_t l_Lean_Name_isAtomic(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__4; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__4; @@ -693,29 +688,34 @@ uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(lean_object*, lean_object*, 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___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_Elab_Term_Quotation_getQuotKind___spec__7(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__10; uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___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*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2; lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__4; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_addNamedQuotInfo___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_828____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1___boxed(lean_object*); lean_object* l_Lean_Elab_Term_Quotation_getAntiquotationIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22; @@ -730,9 +730,11 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ lean_object* l_Std_RBNode_find___at___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24(lean_object*, lean_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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___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_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16119_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16327_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17(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*); @@ -743,16 +745,16 @@ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation_____ uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__43; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__7; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; @@ -760,15 +762,21 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30(lean_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_Quotation_elabNoErrorIfUnused___spec__1___rarg(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2; static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_getQuotKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__7; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1; static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__13___closed__1; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getQuotKind___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_dropRight(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__6___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__7(lean_object*, lean_object*); @@ -777,20 +785,23 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__5; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__20; LEAN_EXPORT lean_object* l_Std_HashSetImp_expand___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__4(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__3; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__4; size_t lean_usize_of_nat(lean_object*); static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable(lean_object*, lean_object*); @@ -798,9 +809,10 @@ uint8_t l_List_foldr___at_List_and___spec__1(uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__2; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate(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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___closed__2; @@ -808,9 +820,10 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__7; static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__13___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__20; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__5___closed__1; @@ -826,7 +839,10 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_addNamedQuotInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; @@ -834,21 +850,17 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__11(lean_object*, size_t, lean_object*, size_t, size_t); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1; lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__29; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__25; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; @@ -875,25 +887,29 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__17; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3; lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts___spec__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_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__6; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -901,49 +917,48 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___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_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_Elab_Term_Quotation_getQuotKind___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_(lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18; lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__3(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Std_HashSetImp_contains___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__22; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__20; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23(lean_object*, 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_expandExplicitBindersAux_loop___spec__1(lean_object*); @@ -957,13 +972,18 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quota static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__10; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); @@ -977,47 +997,50 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ lean_object* l_Lean_Syntax_getAntiquotSpliceSuffix(lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4; lean_object* l_Lean_MacroScopesView_review(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__20; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__11; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2; lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_Quotation_getQuotKind___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*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_elabNoErrorIfUnused___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__8; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(lean_object*, lean_object*, 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_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_exprToSyntax___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14(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*); @@ -1026,22 +1049,23 @@ LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_ static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; lean_object* l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(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*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__22; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2; lean_object* l_Lean_Elab_Term_adaptExpander(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_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__3; lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__19; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; @@ -1049,6 +1073,7 @@ lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13___boxed(lean_object**); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1; 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_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -1079,8 +1104,10 @@ uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27; lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1(lean_object*); @@ -1090,14 +1117,19 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16(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_panic_fn(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_match__syntax_expand___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*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4; lean_object* l_Lean_Syntax_getAntiquotSpliceContents(lean_object*); @@ -1106,6 +1138,7 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ LEAN_EXPORT lean_object* l_List_replace___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__8___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__36; lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4; @@ -1116,11 +1149,10 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabMatchSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__16; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__11; @@ -1130,28 +1162,26 @@ static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_E LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(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_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_adaptRhs(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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__35; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__6; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__8; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__7(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9___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_string_length(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__11; @@ -1160,9 +1190,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Quotation_stxQuot_expand___closed__39; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable_loop(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isEscapedAntiquot(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34; @@ -1171,11 +1201,12 @@ lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__5___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___boxed(lean_object*); lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabNoErrorIfUnused(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19; @@ -1188,7 +1219,9 @@ static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__26; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; @@ -1203,19 +1236,17 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6; lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__31; lean_object* l_Lean_Syntax_antiquotSpliceKind_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__7; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__3___closed__2; @@ -1239,25 +1270,24 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; extern lean_object* l_Lean_Elab_Term_Quotation_hygiene; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_match__syntax_expand___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*); @@ -1266,7 +1296,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__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*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__6___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2; @@ -1274,28 +1304,25 @@ static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__16; lean_object* l_Lean_Syntax_getCanonicalAntiquot(lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__3___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__51; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1; static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__20; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__22; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build(lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__9; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__27; lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_findUsedAlts_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* l_List_appendTR___rarg(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1309,11 +1336,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__1; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused___closed__3; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); @@ -1324,8 +1349,8 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Quotation_0_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__11; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6; LEAN_EXPORT lean_object* l_List_elem___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__9___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1333,17 +1358,21 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__9; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo(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_Quotation_stxQuot_expand___closed__19; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__9___lambda__1___closed__3; @@ -1351,6 +1380,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3; static lean_object* l_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; @@ -1358,7 +1388,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; uint8_t l_Lean_Syntax_isAntiquotSplice(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__1(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_object* x_11) { _start: @@ -4008,6 +4038,520 @@ return x_62; } } } +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__5; +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; 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 = lean_ctor_get(x_1, 1); +x_5 = lean_box(2); +x_6 = l_Lean_Syntax_mkStrLit(x_3, x_5); +x_7 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(x_4); +x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_9 = lean_array_push(x_8, x_6); +x_10 = lean_array_push(x_9, x_7); +x_11 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +x_12 = l_Lean_Syntax_mkCApp(x_11, x_10); +return x_12; +} +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Syntax.Preresolved.namespace", 28); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2; +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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Syntax", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__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_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Preresolved", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("namespace", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__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_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12; +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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("quotedName", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(".", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___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; +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; +x_7 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9; +x_8 = l_Lean_addMacroScope(x_6, x_7, x_5); +x_9 = lean_box(0); +x_10 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3; +x_11 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14; +lean_inc(x_2); +x_12 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_10); +lean_ctor_set(x_12, 2, x_8); +lean_ctor_set(x_12, 3, x_11); +lean_inc(x_1); +x_13 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_9, x_1); +x_14 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_15 = lean_array_push(x_14, x_12); +if (lean_obj_tag(x_13) == 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; +x_16 = l_Lean_quoteNameMk(x_1); +x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_18 = lean_array_push(x_17, x_16); +x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_2); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_18); +x_21 = lean_array_push(x_15, x_20); +x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_2); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, 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_4); +return x_24; +} +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_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_1); +x_25 = lean_ctor_get(x_13, 0); +lean_inc(x_25); +lean_dec(x_13); +x_26 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; +x_27 = l_String_intercalate(x_26, x_25); +x_28 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +x_30 = lean_box(2); +x_31 = l_Lean_Syntax_mkNameLit(x_29, x_30); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_33 = lean_array_push(x_32, x_31); +x_34 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_30); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_33); +x_36 = lean_array_push(x_32, x_35); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_2); +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_2); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = lean_array_push(x_15, x_38); +x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_2); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, 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_4); +return x_42; +} +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Syntax.Preresolved.decl", 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2; +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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decl", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11; +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6; +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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___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_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___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; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___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; lean_object* x_8; 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_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; +x_8 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5; +x_9 = l_Lean_addMacroScope(x_7, x_8, x_6); +x_10 = lean_box(0); +x_11 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3; +x_12 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8; +lean_inc(x_3); +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_9); +lean_ctor_set(x_13, 3, x_12); +lean_inc(x_1); +x_14 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_10, x_1); +x_15 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(x_2); +x_16 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_17 = lean_array_push(x_16, x_13); +if (lean_obj_tag(x_14) == 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; +x_18 = l_Lean_quoteNameMk(x_1); +x_19 = lean_array_push(x_16, x_18); +x_20 = lean_array_push(x_19, x_15); +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_3); +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_3); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_20); +x_23 = lean_array_push(x_17, x_22); +x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_3); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_23); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_5); +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; 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_dec(x_1); +x_27 = lean_ctor_get(x_14, 0); +lean_inc(x_27); +lean_dec(x_14); +x_28 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; +x_29 = l_String_intercalate(x_28, x_27); +x_30 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; +x_31 = lean_string_append(x_30, x_29); +lean_dec(x_29); +x_32 = lean_box(2); +x_33 = l_Lean_Syntax_mkNameLit(x_31, x_32); +x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_35 = lean_array_push(x_34, x_33); +x_36 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_35); +x_38 = lean_array_push(x_16, x_37); +x_39 = lean_array_push(x_38, x_15); +x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_3); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_array_push(x_17, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, 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_5); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous(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; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1), 4, 1); +lean_closure_set(x_3, 0, x_2); +x_4 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_5 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_5, 0, x_4); +lean_closure_set(x_5, 1, x_3); +x_6 = l_Lean_Unhygienic_run___rarg(x_5); +return x_6; +} +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; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___boxed), 5, 2); +lean_closure_set(x_9, 0, x_7); +lean_closure_set(x_9, 1, x_8); +x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_11, 0, x_10); +lean_closure_set(x_11, 1, x_9); +x_12 = l_Lean_Unhygienic_run___rarg(x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___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_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1() { _start: { @@ -5700,40 +6244,6 @@ return x_27; } } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("quotedName", 10); -return x_1; -} -} -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(".", 1); -return x_1; -} -} -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); -return x_1; -} -} LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(lean_object* x_1) { _start: { @@ -5773,16 +6283,16 @@ lean_dec(x_3); x_14 = lean_ctor_get(x_6, 0); lean_inc(x_14); lean_dec(x_6); -x_15 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_15 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_16 = l_String_intercalate(x_15, x_14); -x_17 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_17 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); x_19 = lean_box(2); x_20 = l_Lean_Syntax_mkNameLit(x_18, x_19); x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_22 = lean_array_push(x_21, x_20); -x_23 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_23 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_19); lean_ctor_set(x_24, 1, x_23); @@ -8557,66 +9067,48 @@ return x_3; static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Syntax", 6); -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33() { _start: { lean_object* x_1; @@ -8624,22 +9116,22 @@ x_1 = lean_mk_string_from_bytes("TSyntaxArray.raw", 16); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8647,7 +9139,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36() { _start: { lean_object* x_1; @@ -8655,17 +9147,17 @@ x_1 = lean_mk_string_from_bytes("TSyntaxArray", 12); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38() { _start: { lean_object* x_1; @@ -8673,61 +9165,61 @@ x_1 = lean_mk_string_from_bytes("raw", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; x_3 = lean_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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8737,7 +9229,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45() { _start: { lean_object* x_1; @@ -8745,22 +9237,22 @@ x_1 = lean_mk_string_from_bytes("Option.map", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8768,7 +9260,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8778,31 +9270,31 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51() { _start: { lean_object* x_1; @@ -8810,22 +9302,22 @@ x_1 = lean_mk_string_from_bytes("TSyntax.raw", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8833,7 +9325,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54() { _start: { lean_object* x_1; @@ -8841,71 +9333,71 @@ x_1 = lean_mk_string_from_bytes("TSyntax", 7); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; x_3 = lean_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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8917,19 +9409,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61; x_3 = lean_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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63() { _start: { lean_object* x_1; @@ -8937,22 +9429,22 @@ x_1 = lean_mk_string_from_bytes("x", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8960,17 +9452,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67() { _start: { lean_object* x_1; @@ -8978,22 +9470,22 @@ x_1 = lean_mk_string_from_bytes("Array.empty.push", 16); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9001,7 +9493,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -9013,7 +9505,7 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71() { _start: { lean_object* x_1; lean_object* x_2; @@ -9022,13 +9514,13 @@ x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9036,7 +9528,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -9046,7 +9538,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -9056,24 +9548,24 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -9256,7 +9748,7 @@ lean_ctor_set(x_161, 1, x_160); x_162 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; x_163 = l_Lean_addMacroScope(x_159, x_162, x_154); x_164 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24; -x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; +x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32; lean_inc(x_152); x_166 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_166, 0, x_152); @@ -9313,7 +9805,7 @@ lean_ctor_set(x_187, 1, x_186); x_188 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; x_189 = l_Lean_addMacroScope(x_185, x_188, x_154); x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24; -x_191 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; +x_191 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32; lean_inc(x_152); x_192 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_192, 0, x_152); @@ -9384,10 +9876,10 @@ lean_inc(x_210); x_219 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_219, 0, x_210); lean_ctor_set(x_219, 1, x_218); -x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; +x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; x_221 = l_Lean_addMacroScope(x_217, x_220, x_212); -x_222 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; -x_223 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; +x_222 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; +x_223 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; lean_inc(x_210); x_224 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_224, 0, x_210); @@ -9439,10 +9931,10 @@ lean_inc(x_210); x_243 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_243, 0, x_210); lean_ctor_set(x_243, 1, x_242); -x_244 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; +x_244 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; x_245 = l_Lean_addMacroScope(x_241, x_244, x_212); -x_246 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; -x_247 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; +x_246 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; +x_247 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; lean_inc(x_210); x_248 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_248, 0, x_210); @@ -9517,12 +10009,12 @@ x_275 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_275, 0, x_264); lean_ctor_set(x_275, 1, x_39); lean_ctor_set(x_275, 2, x_274); -x_276 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +x_276 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; lean_inc(x_266); lean_inc(x_271); x_277 = l_Lean_addMacroScope(x_271, x_276, x_266); -x_278 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; -x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +x_278 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; +x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; lean_inc(x_264); x_280 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_280, 0, x_264); @@ -9539,12 +10031,12 @@ lean_inc(x_264); x_284 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_284, 0, x_264); lean_ctor_set(x_284, 1, x_283); -x_285 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_285 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; lean_inc(x_266); lean_inc(x_271); x_286 = l_Lean_addMacroScope(x_271, x_285, x_266); -x_287 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; -x_288 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_287 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_288 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; lean_inc(x_264); x_289 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_289, 0, x_264); @@ -9616,7 +10108,7 @@ lean_ctor_set(x_319, 2, x_318); lean_inc(x_275); x_320 = lean_array_push(x_290, x_275); x_321 = lean_array_push(x_320, x_319); -x_322 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_322 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_inc(x_264); x_323 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_323, 0, x_264); @@ -9643,18 +10135,18 @@ lean_inc(x_266); lean_inc(x_271); x_331 = l_Lean_addMacroScope(x_271, x_330, x_266); x_332 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__3; -x_333 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; +x_333 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; lean_inc(x_264); x_334 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_334, 0, x_264); lean_ctor_set(x_334, 1, x_332); lean_ctor_set(x_334, 2, x_331); lean_ctor_set(x_334, 3, x_333); -x_335 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; +x_335 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; lean_inc(x_266); lean_inc(x_271); x_336 = l_Lean_addMacroScope(x_271, x_335, x_266); -x_337 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; +x_337 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; lean_inc(x_264); x_338 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_338, 0, x_264); @@ -9701,7 +10193,7 @@ x_353 = l_Lean_Name_str___override(x_351, x_352); lean_inc(x_266); lean_inc(x_271); x_354 = l_Lean_addMacroScope(x_271, x_353, x_266); -x_355 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +x_355 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; lean_inc(x_351); x_356 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_356, 0, x_351); @@ -9709,7 +10201,7 @@ lean_ctor_set(x_356, 1, x_355); x_357 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_357, 0, x_356); lean_ctor_set(x_357, 1, x_123); -x_358 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; +x_358 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; lean_inc(x_264); x_359 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_359, 0, x_264); @@ -9736,12 +10228,12 @@ x_369 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_369, 0, x_264); lean_ctor_set(x_369, 1, x_368); lean_ctor_set(x_369, 2, x_367); -x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; lean_inc(x_266); lean_inc(x_271); x_371 = l_Lean_addMacroScope(x_271, x_370, x_266); -x_372 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; -x_373 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78; +x_372 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +x_373 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; lean_inc(x_264); x_374 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_374, 0, x_264); @@ -9837,12 +10329,12 @@ x_410 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_410, 0, x_264); lean_ctor_set(x_410, 1, x_39); lean_ctor_set(x_410, 2, x_409); -x_411 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +x_411 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; lean_inc(x_266); lean_inc(x_406); x_412 = l_Lean_addMacroScope(x_406, x_411, x_266); -x_413 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; -x_414 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +x_413 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; +x_414 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; lean_inc(x_264); x_415 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_415, 0, x_264); @@ -9859,12 +10351,12 @@ lean_inc(x_264); x_419 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_419, 0, x_264); lean_ctor_set(x_419, 1, x_418); -x_420 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_420 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; lean_inc(x_266); lean_inc(x_406); x_421 = l_Lean_addMacroScope(x_406, x_420, x_266); -x_422 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; -x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_422 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; lean_inc(x_264); x_424 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_424, 0, x_264); @@ -9936,7 +10428,7 @@ lean_ctor_set(x_454, 2, x_453); lean_inc(x_410); x_455 = lean_array_push(x_425, x_410); x_456 = lean_array_push(x_455, x_454); -x_457 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_457 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_inc(x_264); x_458 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_458, 0, x_264); @@ -9963,18 +10455,18 @@ lean_inc(x_266); lean_inc(x_406); x_466 = l_Lean_addMacroScope(x_406, x_465, x_266); x_467 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__3; -x_468 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; +x_468 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; lean_inc(x_264); x_469 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_469, 0, x_264); lean_ctor_set(x_469, 1, x_467); lean_ctor_set(x_469, 2, x_466); lean_ctor_set(x_469, 3, x_468); -x_470 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; +x_470 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; lean_inc(x_266); lean_inc(x_406); x_471 = l_Lean_addMacroScope(x_406, x_470, x_266); -x_472 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; +x_472 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; lean_inc(x_264); x_473 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_473, 0, x_264); @@ -10021,7 +10513,7 @@ x_488 = l_Lean_Name_str___override(x_486, x_487); lean_inc(x_266); lean_inc(x_406); x_489 = l_Lean_addMacroScope(x_406, x_488, x_266); -x_490 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +x_490 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; lean_inc(x_486); x_491 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_491, 0, x_486); @@ -10029,7 +10521,7 @@ lean_ctor_set(x_491, 1, x_490); x_492 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_492, 0, x_491); lean_ctor_set(x_492, 1, x_123); -x_493 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; +x_493 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; lean_inc(x_264); x_494 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_494, 0, x_264); @@ -10056,12 +10548,12 @@ x_504 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_504, 0, x_264); lean_ctor_set(x_504, 1, x_503); lean_ctor_set(x_504, 2, x_502); -x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; lean_inc(x_266); lean_inc(x_406); x_506 = l_Lean_addMacroScope(x_406, x_505, x_266); -x_507 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; -x_508 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78; +x_507 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +x_508 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; lean_inc(x_264); x_509 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_509, 0, x_264); @@ -10870,6 +11362,243 @@ return x_27; } } } +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(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 +{ +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 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_9); +{ +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_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_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_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_2); +x_1 = x_12; +x_2 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___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; +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; +x_8 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9; +x_9 = l_Lean_addMacroScope(x_7, x_8, x_6); +x_10 = lean_box(0); +x_11 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3; +x_12 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14; +lean_inc(x_3); +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_9); +lean_ctor_set(x_13, 3, x_12); +lean_inc(x_1); +x_14 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_10, x_1); +x_15 = lean_array_push(x_2, x_13); +if (lean_obj_tag(x_14) == 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; +x_16 = l_Lean_quoteNameMk(x_1); +x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_18 = lean_array_push(x_17, x_16); +x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_3); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_18); +x_21 = lean_array_push(x_15, x_20); +x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_3); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, 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_5); +return x_24; +} +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_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_1); +x_25 = lean_ctor_get(x_14, 0); +lean_inc(x_25); +lean_dec(x_14); +x_26 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; +x_27 = l_String_intercalate(x_26, x_25); +x_28 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +x_30 = lean_box(2); +x_31 = l_Lean_Syntax_mkNameLit(x_29, x_30); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_33 = lean_array_push(x_32, x_31); +x_34 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_30); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_33); +x_36 = lean_array_push(x_32, x_35); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_3); +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_3); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = lean_array_push(x_15, x_38); +x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, 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_5); +return x_42; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___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; lean_object* x_8; 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_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; +x_9 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = lean_box(0); +x_12 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3; +x_13 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8; +lean_inc(x_4); +x_14 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_12); +lean_ctor_set(x_14, 2, x_10); +lean_ctor_set(x_14, 3, x_13); +lean_inc(x_1); +x_15 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_11, x_1); +x_16 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(x_2); +lean_inc(x_3); +x_17 = lean_array_push(x_3, x_14); +if (lean_obj_tag(x_15) == 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; +x_18 = l_Lean_quoteNameMk(x_1); +x_19 = lean_array_push(x_3, x_18); +x_20 = lean_array_push(x_19, x_16); +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_4); +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_4); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_20); +x_23 = lean_array_push(x_17, x_22); +x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_4); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_23); +x_26 = lean_alloc_ctor(0, 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; 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_dec(x_1); +x_27 = lean_ctor_get(x_15, 0); +lean_inc(x_27); +lean_dec(x_15); +x_28 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; +x_29 = l_String_intercalate(x_28, x_27); +x_30 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; +x_31 = lean_string_append(x_30, x_29); +lean_dec(x_29); +x_32 = lean_box(2); +x_33 = l_Lean_Syntax_mkNameLit(x_31, x_32); +x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_35 = lean_array_push(x_34, x_33); +x_36 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_35); +x_38 = lean_array_push(x_3, x_37); +x_39 = lean_array_push(x_38, x_16); +x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_4); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_4); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_array_push(x_17, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_4); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, 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_6); +return x_45; +} +} +} LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(lean_object* x_1) { _start: { @@ -10881,22 +11610,62 @@ return x_2; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +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); -x_5 = lean_box(2); -x_6 = l_Lean_Syntax_mkStrLit(x_3, x_5); -x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_4); -x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_9 = lean_array_push(x_8, x_6); -x_10 = lean_array_push(x_9, x_7); -x_11 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; -x_12 = l_Lean_Syntax_mkCApp(x_11, x_10); -return x_12; +lean_inc(x_4); +lean_dec(x_1); +x_5 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_4); +if (lean_obj_tag(x_3) == 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; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +lean_dec(x_3); +x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_8 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__1), 5, 2); +lean_closure_set(x_8, 0, x_6); +lean_closure_set(x_8, 1, x_7); +x_9 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_8); +x_11 = l_Lean_Unhygienic_run___rarg(x_10); +x_12 = lean_array_push(x_7, x_11); +x_13 = lean_array_push(x_12, x_5); +x_14 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +x_15 = l_Lean_Syntax_mkCApp(x_14, x_13); +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; lean_object* x_25; lean_object* x_26; +x_16 = lean_ctor_get(x_3, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_dec(x_3); +x_18 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_19 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__2___boxed), 6, 3); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_17); +lean_closure_set(x_19, 2, x_18); +x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_21 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_21, 0, x_20); +lean_closure_set(x_21, 1, x_19); +x_22 = l_Lean_Unhygienic_run___rarg(x_21); +x_23 = lean_array_push(x_18, x_22); +x_24 = lean_array_push(x_23, x_5); +x_25 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +x_26 = l_Lean_Syntax_mkCApp(x_25, x_24); +return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10913,7 +11682,7 @@ lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_4); +x_5 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(x_4); x_6 = lean_ctor_get(x_3, 0); lean_inc(x_6); x_7 = lean_ctor_get(x_3, 1); @@ -10922,7 +11691,7 @@ lean_dec(x_3); x_8 = lean_box(0); lean_inc(x_6); x_9 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_8, x_6); -x_10 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_7); +x_10 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___spec__1(x_7); lean_dec(x_7); if (lean_obj_tag(x_9) == 0) { @@ -10946,16 +11715,16 @@ lean_dec(x_6); x_21 = lean_ctor_get(x_9, 0); lean_inc(x_21); lean_dec(x_9); -x_22 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_22 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_23 = l_String_intercalate(x_22, x_21); -x_24 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_24 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_25 = lean_string_append(x_24, x_23); lean_dec(x_23); x_26 = lean_box(2); x_27 = l_Lean_Syntax_mkNameLit(x_25, x_26); x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_29 = lean_array_push(x_28, x_27); -x_30 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_30 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_26); lean_ctor_set(x_31, 1, x_30); @@ -11026,66 +11795,56 @@ return x_4; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("node", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8; 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___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_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -11093,22 +11852,22 @@ x_1 = lean_mk_string_from_bytes("info", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11116,17 +11875,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -11134,17 +11893,17 @@ x_1 = lean_mk_string_from_bytes("appendCore", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__2; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -11152,16 +11911,16 @@ x_1 = lean_mk_string_from_bytes("unexpected antiquotation splice", 31); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -11169,16 +11928,16 @@ x_1 = lean_mk_string_from_bytes("expected token", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -11186,22 +11945,22 @@ x_1 = lean_mk_string_from_bytes("Syntax.atom", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11209,7 +11968,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -11217,51 +11976,51 @@ x_1 = lean_mk_string_from_bytes("atom", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27; x_3 = lean_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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -11269,22 +12028,22 @@ x_1 = lean_mk_string_from_bytes("SourceInfo.fromRef", 18); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11292,7 +12051,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -11300,17 +12059,17 @@ x_1 = lean_mk_string_from_bytes("SourceInfo", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34() { _start: { lean_object* x_1; @@ -11318,54 +12077,54 @@ x_1 = lean_mk_string_from_bytes("fromRef", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__33; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__34; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37; 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -11454,10 +12213,10 @@ lean_inc(x_328); x_338 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_338, 0, x_328); lean_ctor_set(x_338, 1, x_337); -x_339 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_339 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; x_340 = l_Lean_addMacroScope(x_335, x_339, x_330); -x_341 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; -x_342 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_341 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_342 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; lean_inc(x_328); x_343 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_343, 0, x_328); @@ -11513,10 +12272,10 @@ lean_inc(x_328); x_366 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_366, 0, x_328); lean_ctor_set(x_366, 1, x_365); -x_367 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_367 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; x_368 = l_Lean_addMacroScope(x_363, x_367, x_330); -x_369 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; -x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_369 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; +x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; lean_inc(x_328); x_371 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_371, 0, x_328); @@ -11616,7 +12375,7 @@ x_22 = 0; if (x_18 == 0) { lean_object* x_133; -x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; +x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16; x_23 = x_133; goto block_132; } @@ -11664,22 +12423,22 @@ x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); x_37 = lean_environment_main_module(x_36); -x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8; +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7; lean_inc(x_32); lean_inc(x_37); x_39 = l_Lean_addMacroScope(x_37, x_38, x_32); x_40 = lean_box(0); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10; lean_inc(x_30); x_43 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_43, 0, x_30); lean_ctor_set(x_43, 1, x_41); lean_ctor_set(x_43, 2, x_39); lean_ctor_set(x_43, 3, x_42); -x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_45 = l_Lean_addMacroScope(x_37, x_44, x_32); -x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_30); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_30); @@ -11721,16 +12480,16 @@ lean_dec(x_2); x_62 = lean_ctor_get(x_48, 0); lean_inc(x_62); lean_dec(x_48); -x_63 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_63 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_64 = l_String_intercalate(x_63, x_62); -x_65 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_65 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_66 = lean_string_append(x_65, x_64); lean_dec(x_64); x_67 = lean_box(2); x_68 = l_Lean_Syntax_mkNameLit(x_66, x_67); x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_70 = lean_array_push(x_69, x_68); -x_71 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_71 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_67); lean_ctor_set(x_72, 1, x_71); @@ -11765,22 +12524,22 @@ x_82 = lean_ctor_get(x_80, 0); lean_inc(x_82); lean_dec(x_80); x_83 = lean_environment_main_module(x_82); -x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8; +x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7; lean_inc(x_32); lean_inc(x_83); x_85 = l_Lean_addMacroScope(x_83, x_84, x_32); x_86 = lean_box(0); x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; -x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; +x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10; lean_inc(x_30); x_89 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_89, 0, x_30); lean_ctor_set(x_89, 1, x_87); lean_ctor_set(x_89, 2, x_85); lean_ctor_set(x_89, 3, x_88); -x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_91 = l_Lean_addMacroScope(x_83, x_90, x_32); -x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_30); x_93 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_93, 0, x_30); @@ -11824,16 +12583,16 @@ lean_dec(x_2); x_109 = lean_ctor_get(x_94, 0); lean_inc(x_109); lean_dec(x_94); -x_110 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_110 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_111 = l_String_intercalate(x_110, x_109); -x_112 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_112 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_113 = lean_string_append(x_112, x_111); lean_dec(x_111); x_114 = lean_box(2); x_115 = l_Lean_Syntax_mkNameLit(x_113, x_114); x_116 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_117 = lean_array_push(x_116, x_115); -x_118 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_118 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_119 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_119, 0, x_114); lean_ctor_set(x_119, 1, x_118); @@ -11910,7 +12669,7 @@ if (x_139 == 0) { lean_object* x_140; lean_object* x_141; lean_dec(x_2); -x_140 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_140 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_141 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_1, x_140, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_141; } @@ -11935,7 +12694,7 @@ if (x_146 == 0) { lean_object* x_147; lean_object* x_148; lean_dec(x_2); -x_147 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_147 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_148 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_1, x_147, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_148; } @@ -11958,7 +12717,7 @@ if (x_151 == 0) { lean_object* x_152; lean_object* x_153; lean_dec(x_2); -x_152 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_152 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_153 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_1, x_152, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_153; } @@ -12014,12 +12773,12 @@ x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); lean_dec(x_168); x_170 = lean_environment_main_module(x_169); -x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; +x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_inc(x_165); lean_inc(x_170); x_172 = l_Lean_addMacroScope(x_170, x_171, x_165); -x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; -x_174 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; +x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_174 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; lean_inc(x_163); x_175 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_175, 0, x_163); @@ -12030,10 +12789,10 @@ x_176 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab lean_inc(x_163); lean_ctor_set(x_158, 1, x_176); lean_ctor_set(x_158, 0, x_163); -x_177 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36; +x_177 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35; x_178 = l_Lean_addMacroScope(x_170, x_177, x_165); -x_179 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; -x_180 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40; +x_179 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; +x_180 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39; lean_inc(x_163); x_181 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_181, 0, x_163); @@ -12118,12 +12877,12 @@ x_215 = lean_ctor_get(x_213, 0); lean_inc(x_215); lean_dec(x_213); x_216 = lean_environment_main_module(x_215); -x_217 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; +x_217 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_inc(x_165); lean_inc(x_216); x_218 = l_Lean_addMacroScope(x_216, x_217, x_165); -x_219 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; -x_220 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; +x_219 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_220 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; lean_inc(x_163); x_221 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_221, 0, x_163); @@ -12134,10 +12893,10 @@ x_222 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab lean_inc(x_163); lean_ctor_set(x_158, 1, x_222); lean_ctor_set(x_158, 0, x_163); -x_223 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36; +x_223 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35; x_224 = l_Lean_addMacroScope(x_216, x_223, x_165); -x_225 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; -x_226 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40; +x_225 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; +x_226 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39; lean_inc(x_163); x_227 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_227, 0, x_163); @@ -12247,12 +13006,12 @@ x_269 = lean_ctor_get(x_266, 0); lean_inc(x_269); lean_dec(x_266); x_270 = lean_environment_main_module(x_269); -x_271 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; +x_271 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_inc(x_264); lean_inc(x_270); x_272 = l_Lean_addMacroScope(x_270, x_271, x_264); -x_273 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; -x_274 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; +x_273 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_274 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; lean_inc(x_262); x_275 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_275, 0, x_262); @@ -12264,10 +13023,10 @@ lean_inc(x_262); x_277 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_277, 0, x_262); lean_ctor_set(x_277, 1, x_276); -x_278 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__36; +x_278 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__35; x_279 = l_Lean_addMacroScope(x_270, x_278, x_264); -x_280 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; -x_281 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40; +x_280 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; +x_281 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39; lean_inc(x_262); x_282 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_282, 0, x_262); @@ -12351,7 +13110,7 @@ else { lean_object* x_315; lean_object* x_316; lean_dec(x_158); -x_315 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21; +x_315 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20; x_316 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_1, x_315, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_316; } @@ -12401,7 +13160,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -12411,7 +13170,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -12637,7 +13396,7 @@ return x_3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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; 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_170; lean_object* x_171; lean_object* x_172; +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_170; lean_object* x_171; lean_inc(x_9); lean_inc(x_1); x_12 = l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveName___spec__1(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); @@ -12651,330 +13410,329 @@ lean_inc(x_1); x_16 = l_Lean_Elab_Term_Quotation_resolveSectionVariable(x_15, x_1); x_17 = l_List_appendTR___rarg(x_13, x_16); x_18 = l_List_appendTR___rarg(x_17, x_2); -x_170 = lean_box(0); +x_19 = lean_box(0); +x_20 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_18, x_19); lean_inc(x_1); -x_171 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_170, x_1); +x_170 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_1); lean_inc(x_9); -x_172 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_exprToSyntax___spec__1___rarg(x_9, x_10, x_14); -if (lean_obj_tag(x_171) == 0) +x_171 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_exprToSyntax___spec__1___rarg(x_9, x_10, x_14); +if (lean_obj_tag(x_170) == 0) { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_172, 0); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = l_Lean_quoteNameMk(x_1); -x_19 = x_175; -x_20 = x_173; +lean_dec(x_171); +x_174 = l_Lean_quoteNameMk(x_1); x_21 = x_174; +x_22 = x_172; +x_23 = x_173; goto block_169; } 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_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_dec(x_1); +x_175 = lean_ctor_get(x_170, 0); +lean_inc(x_175); +lean_dec(x_170); x_176 = lean_ctor_get(x_171, 0); lean_inc(x_176); -lean_dec(x_171); -x_177 = lean_ctor_get(x_172, 0); +x_177 = lean_ctor_get(x_171, 1); lean_inc(x_177); -x_178 = lean_ctor_get(x_172, 1); -lean_inc(x_178); -lean_dec(x_172); -x_179 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; -x_180 = l_String_intercalate(x_179, x_176); -x_181 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; -x_182 = lean_string_append(x_181, x_180); -lean_dec(x_180); -x_183 = lean_box(2); -x_184 = l_Lean_Syntax_mkNameLit(x_182, x_183); -x_185 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_186 = lean_array_push(x_185, x_184); -x_187 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; -x_188 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_188, 0, x_183); -lean_ctor_set(x_188, 1, x_187); -lean_ctor_set(x_188, 2, x_186); -x_19 = x_188; -x_20 = x_177; -x_21 = x_178; +lean_dec(x_171); +x_178 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; +x_179 = l_String_intercalate(x_178, x_175); +x_180 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; +x_181 = lean_string_append(x_180, x_179); +lean_dec(x_179); +x_182 = lean_box(2); +x_183 = l_Lean_Syntax_mkNameLit(x_181, x_182); +x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_185 = lean_array_push(x_184, x_183); +x_186 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; +x_187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_187, 0, x_182); +lean_ctor_set(x_187, 1, x_186); +lean_ctor_set(x_187, 2, x_185); +x_21 = x_187; +x_22 = x_176; +x_23 = x_177; goto block_169; } block_169: { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_9, 10); -lean_inc(x_22); +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_9, 10); +lean_inc(x_24); lean_dec(x_9); -x_23 = lean_st_ref_get(x_10, x_21); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_25 = lean_st_ref_get(x_10, x_23); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -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; -x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_environment_main_module(x_26); -x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__5; +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; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_environment_main_module(x_28); +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__5; +lean_inc(x_24); +lean_inc(x_29); +x_31 = l_Lean_addMacroScope(x_29, x_30, x_24); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; lean_inc(x_22); -lean_inc(x_27); -x_29 = l_Lean_addMacroScope(x_27, x_28, x_22); -x_30 = lean_box(0); -x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; -lean_inc(x_20); -x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_31); -lean_ctor_set(x_33, 2, x_29); -lean_ctor_set(x_33, 3, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_22); +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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +lean_inc(x_24); +lean_inc(x_29); +x_36 = l_Lean_addMacroScope(x_29, x_35, x_24); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_22); -lean_inc(x_27); -x_35 = l_Lean_addMacroScope(x_27, x_34, x_22); -x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; -lean_inc(x_20); -x_37 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_37, 0, x_20); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_35); -lean_ctor_set(x_37, 3, x_30); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; -lean_inc(x_20); -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_20); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__16; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_19); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_22); -lean_inc(x_27); -x_41 = l_Lean_addMacroScope(x_27, x_40, x_22); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__15; -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__19; -lean_inc(x_20); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_20); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_43); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__23; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__16; +lean_inc(x_24); +lean_inc(x_29); +x_42 = l_Lean_addMacroScope(x_29, x_41, x_24); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__15; +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__19; lean_inc(x_22); -lean_inc(x_27); -x_46 = l_Lean_addMacroScope(x_27, x_45, x_22); -x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__22; -lean_inc(x_20); -x_48 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_48, 0, x_20); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_46); -lean_ctor_set(x_48, 3, x_30); -x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__27; -x_50 = l_Lean_addMacroScope(x_27, x_49, x_22); -x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__26; -lean_inc(x_20); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_20); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_50); -lean_ctor_set(x_52, 3, x_30); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; -x_54 = lean_array_push(x_53, x_48); -x_55 = lean_array_push(x_54, x_19); -x_56 = lean_array_push(x_55, x_52); -x_57 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_20); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_20); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_60 = lean_array_push(x_59, x_44); -x_61 = lean_array_push(x_60, x_58); -x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -lean_inc(x_20); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_20); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_61); -x_64 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_20); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_20); -lean_ctor_set(x_65, 1, x_57); -lean_ctor_set(x_65, 2, x_64); -x_66 = lean_array_push(x_59, x_63); -x_67 = lean_array_push(x_66, x_65); -lean_inc(x_20); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_20); -lean_ctor_set(x_68, 1, x_57); -lean_ctor_set(x_68, 2, x_67); -x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; -lean_inc(x_20); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_20); -lean_ctor_set(x_70, 1, x_69); -x_71 = lean_array_push(x_53, x_39); -x_72 = lean_array_push(x_71, x_68); -x_73 = lean_array_push(x_72, x_70); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; -lean_inc(x_20); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_20); -lean_ctor_set(x_75, 1, x_74); -lean_ctor_set(x_75, 2, x_73); -x_76 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_18); -x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_78 = lean_array_push(x_77, x_37); -x_79 = lean_array_push(x_59, x_33); -x_80 = lean_ctor_get(x_3, 0); -x_81 = lean_ctor_get(x_3, 1); -x_82 = lean_ctor_get(x_3, 2); -x_83 = lean_string_utf8_extract(x_80, x_81, x_82); -x_84 = lean_box(2); -x_85 = l_Lean_Syntax_mkStrLit(x_83, x_84); -lean_dec(x_83); -x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_87 = lean_array_push(x_86, x_85); -x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__12; -x_89 = l_Lean_Syntax_mkCApp(x_88, x_87); -x_90 = lean_array_push(x_78, x_89); -x_91 = lean_array_push(x_90, x_75); +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_42); +lean_ctor_set(x_45, 3, x_44); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__23; +lean_inc(x_24); +lean_inc(x_29); +x_47 = l_Lean_addMacroScope(x_29, x_46, x_24); +x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__22; +lean_inc(x_22); +x_49 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_49, 0, x_22); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_47); +lean_ctor_set(x_49, 3, x_19); +x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__27; +x_51 = l_Lean_addMacroScope(x_29, x_50, x_24); +x_52 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__26; +lean_inc(x_22); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_22); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +lean_ctor_set(x_53, 3, x_19); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_55 = lean_array_push(x_54, x_49); +x_56 = lean_array_push(x_55, x_21); +x_57 = lean_array_push(x_56, x_53); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_22); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_22); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_61 = lean_array_push(x_60, x_45); +x_62 = lean_array_push(x_61, x_59); +x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_inc(x_22); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_22); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_62); +x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_22); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_22); +lean_ctor_set(x_66, 1, x_58); +lean_ctor_set(x_66, 2, x_65); +x_67 = lean_array_push(x_60, x_64); +x_68 = lean_array_push(x_67, x_66); +lean_inc(x_22); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_22); +lean_ctor_set(x_69, 1, x_58); +lean_ctor_set(x_69, 2, x_68); +x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +lean_inc(x_22); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_22); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_array_push(x_54, x_40); +x_73 = lean_array_push(x_72, x_69); +x_74 = lean_array_push(x_73, x_71); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +lean_inc(x_22); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_22); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_20); +x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_79 = lean_array_push(x_78, x_38); +x_80 = lean_array_push(x_60, x_34); +x_81 = lean_ctor_get(x_3, 0); +x_82 = lean_ctor_get(x_3, 1); +x_83 = lean_ctor_get(x_3, 2); +x_84 = lean_string_utf8_extract(x_81, x_82, x_83); +x_85 = lean_box(2); +x_86 = l_Lean_Syntax_mkStrLit(x_84, x_85); +lean_dec(x_84); +x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_88 = lean_array_push(x_87, x_86); +x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__12; +x_90 = l_Lean_Syntax_mkCApp(x_89, x_88); +x_91 = lean_array_push(x_79, x_90); x_92 = lean_array_push(x_91, x_76); -lean_inc(x_20); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_20); -lean_ctor_set(x_93, 1, x_57); -lean_ctor_set(x_93, 2, x_92); -x_94 = lean_array_push(x_79, x_93); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_20); -lean_ctor_set(x_95, 1, x_62); -lean_ctor_set(x_95, 2, x_94); -lean_ctor_set(x_23, 0, x_95); -return x_23; +x_93 = lean_array_push(x_92, x_77); +lean_inc(x_22); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_22); +lean_ctor_set(x_94, 1, x_58); +lean_ctor_set(x_94, 2, x_93); +x_95 = lean_array_push(x_80, x_94); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_22); +lean_ctor_set(x_96, 1, x_63); +lean_ctor_set(x_96, 2, x_95); +lean_ctor_set(x_25, 0, x_96); +return x_25; } 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; 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; -x_96 = lean_ctor_get(x_23, 0); -x_97 = lean_ctor_get(x_23, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_23); -x_98 = lean_ctor_get(x_96, 0); +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; +x_97 = lean_ctor_get(x_25, 0); +x_98 = lean_ctor_get(x_25, 1); lean_inc(x_98); -lean_dec(x_96); -x_99 = lean_environment_main_module(x_98); -x_100 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__5; -lean_inc(x_22); +lean_inc(x_97); +lean_dec(x_25); +x_99 = lean_ctor_get(x_97, 0); lean_inc(x_99); -x_101 = l_Lean_addMacroScope(x_99, x_100, x_22); -x_102 = lean_box(0); +lean_dec(x_97); +x_100 = lean_environment_main_module(x_99); +x_101 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__5; +lean_inc(x_24); +lean_inc(x_100); +x_102 = l_Lean_addMacroScope(x_100, x_101, x_24); x_103 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; x_104 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; -lean_inc(x_20); +lean_inc(x_22); x_105 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_105, 0, x_20); +lean_ctor_set(x_105, 0, x_22); lean_ctor_set(x_105, 1, x_103); -lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 2, x_102); lean_ctor_set(x_105, 3, x_104); -x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +lean_inc(x_24); +lean_inc(x_100); +x_107 = l_Lean_addMacroScope(x_100, x_106, x_24); +x_108 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_22); -lean_inc(x_99); -x_107 = l_Lean_addMacroScope(x_99, x_106, x_22); -x_108 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; -lean_inc(x_20); x_109 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_109, 0, x_20); +lean_ctor_set(x_109, 0, x_22); lean_ctor_set(x_109, 1, x_108); lean_ctor_set(x_109, 2, x_107); -lean_ctor_set(x_109, 3, x_102); +lean_ctor_set(x_109, 3, x_19); x_110 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; -lean_inc(x_20); +lean_inc(x_22); x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_20); +lean_ctor_set(x_111, 0, x_22); lean_ctor_set(x_111, 1, x_110); x_112 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__16; -lean_inc(x_22); -lean_inc(x_99); -x_113 = l_Lean_addMacroScope(x_99, x_112, x_22); +lean_inc(x_24); +lean_inc(x_100); +x_113 = l_Lean_addMacroScope(x_100, x_112, x_24); x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__15; x_115 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__19; -lean_inc(x_20); +lean_inc(x_22); x_116 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_116, 0, x_20); +lean_ctor_set(x_116, 0, x_22); lean_ctor_set(x_116, 1, x_114); lean_ctor_set(x_116, 2, x_113); lean_ctor_set(x_116, 3, x_115); x_117 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__23; -lean_inc(x_22); -lean_inc(x_99); -x_118 = l_Lean_addMacroScope(x_99, x_117, x_22); +lean_inc(x_24); +lean_inc(x_100); +x_118 = l_Lean_addMacroScope(x_100, x_117, x_24); x_119 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__22; -lean_inc(x_20); +lean_inc(x_22); x_120 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_120, 0, x_20); +lean_ctor_set(x_120, 0, x_22); lean_ctor_set(x_120, 1, x_119); lean_ctor_set(x_120, 2, x_118); -lean_ctor_set(x_120, 3, x_102); +lean_ctor_set(x_120, 3, x_19); x_121 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__27; -x_122 = l_Lean_addMacroScope(x_99, x_121, x_22); +x_122 = l_Lean_addMacroScope(x_100, x_121, x_24); x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__26; -lean_inc(x_20); +lean_inc(x_22); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_20); +lean_ctor_set(x_124, 0, x_22); lean_ctor_set(x_124, 1, x_123); lean_ctor_set(x_124, 2, x_122); -lean_ctor_set(x_124, 3, x_102); +lean_ctor_set(x_124, 3, x_19); x_125 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_126 = lean_array_push(x_125, x_120); -x_127 = lean_array_push(x_126, x_19); +x_127 = lean_array_push(x_126, x_21); x_128 = lean_array_push(x_127, x_124); x_129 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_20); +lean_inc(x_22); x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_20); +lean_ctor_set(x_130, 0, x_22); lean_ctor_set(x_130, 1, x_129); lean_ctor_set(x_130, 2, x_128); x_131 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_132 = lean_array_push(x_131, x_116); x_133 = lean_array_push(x_132, x_130); x_134 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -lean_inc(x_20); +lean_inc(x_22); x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_20); +lean_ctor_set(x_135, 0, x_22); lean_ctor_set(x_135, 1, x_134); lean_ctor_set(x_135, 2, x_133); x_136 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_20); +lean_inc(x_22); x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_20); +lean_ctor_set(x_137, 0, x_22); lean_ctor_set(x_137, 1, x_129); lean_ctor_set(x_137, 2, x_136); x_138 = lean_array_push(x_131, x_135); x_139 = lean_array_push(x_138, x_137); -lean_inc(x_20); +lean_inc(x_22); x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_20); +lean_ctor_set(x_140, 0, x_22); lean_ctor_set(x_140, 1, x_129); lean_ctor_set(x_140, 2, x_139); x_141 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; -lean_inc(x_20); +lean_inc(x_22); x_142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_142, 0, x_20); +lean_ctor_set(x_142, 0, x_22); lean_ctor_set(x_142, 1, x_141); x_143 = lean_array_push(x_125, x_111); x_144 = lean_array_push(x_143, x_140); x_145 = lean_array_push(x_144, x_142); x_146 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; -lean_inc(x_20); +lean_inc(x_22); x_147 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_147, 0, x_20); +lean_ctor_set(x_147, 0, x_22); lean_ctor_set(x_147, 1, x_146); lean_ctor_set(x_147, 2, x_145); -x_148 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_18); +x_148 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_20); x_149 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_150 = lean_array_push(x_149, x_109); x_151 = lean_array_push(x_131, x_105); @@ -12992,19 +13750,19 @@ x_161 = l_Lean_Syntax_mkCApp(x_160, x_159); x_162 = lean_array_push(x_150, x_161); x_163 = lean_array_push(x_162, x_147); x_164 = lean_array_push(x_163, x_148); -lean_inc(x_20); +lean_inc(x_22); x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_20); +lean_ctor_set(x_165, 0, x_22); lean_ctor_set(x_165, 1, x_129); lean_ctor_set(x_165, 2, x_164); x_166 = lean_array_push(x_151, x_165); x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_20); +lean_ctor_set(x_167, 0, x_22); lean_ctor_set(x_167, 1, x_134); lean_ctor_set(x_167, 2, 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_97); +lean_ctor_set(x_168, 1, x_98); return x_168; } } @@ -13018,6 +13776,30 @@ x_1 = l_Lean_Elab_Term_Quotation_hygiene; return x_1; } } +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__6; +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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax(lean_object* x_1, lean_object* x_2, 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: { @@ -13153,22 +13935,22 @@ x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); x_37 = lean_environment_main_module(x_36); -x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_inc(x_32); lean_inc(x_37); x_39 = l_Lean_addMacroScope(x_37, x_38, x_32); x_40 = lean_box(0); -x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; lean_inc(x_30); x_43 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_43, 0, x_30); lean_ctor_set(x_43, 1, x_41); lean_ctor_set(x_43, 2, x_39); lean_ctor_set(x_43, 3, x_42); -x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_45 = l_Lean_addMacroScope(x_37, x_44, x_32); -x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_30); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_30); @@ -13209,22 +13991,22 @@ x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); lean_dec(x_59); x_62 = lean_environment_main_module(x_61); -x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; +x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_inc(x_32); lean_inc(x_62); x_64 = l_Lean_addMacroScope(x_62, x_63, x_32); x_65 = lean_box(0); -x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; -x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; +x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; +x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; lean_inc(x_30); x_68 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_68, 0, x_30); lean_ctor_set(x_68, 1, x_66); lean_ctor_set(x_68, 2, x_64); lean_ctor_set(x_68, 3, x_67); -x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_70 = l_Lean_addMacroScope(x_62, x_69, x_32); -x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_30); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_30); @@ -13307,15 +14089,15 @@ lean_inc(x_101); x_103 = l_Lean_addMacroScope(x_101, x_102, x_96); x_104 = lean_box(0); x_105 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; -x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; +x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; lean_inc(x_94); lean_ctor_set(x_1, 3, x_106); lean_ctor_set(x_1, 2, x_103); lean_ctor_set(x_1, 1, x_105); lean_ctor_set(x_1, 0, x_94); -x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_108 = l_Lean_addMacroScope(x_101, x_107, x_96); -x_109 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_109 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_94); x_110 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_110, 0, x_94); @@ -13324,7 +14106,7 @@ lean_ctor_set(x_110, 2, x_108); lean_ctor_set(x_110, 3, x_104); lean_inc(x_87); x_111 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_104, x_87); -x_112 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_88); +x_112 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(x_88); x_113 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_114 = lean_array_push(x_113, x_110); x_115 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; @@ -13376,14 +14158,14 @@ lean_dec(x_87); x_136 = lean_ctor_get(x_111, 0); lean_inc(x_136); lean_dec(x_111); -x_137 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_137 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_138 = l_String_intercalate(x_137, x_136); -x_139 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_139 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_140 = lean_string_append(x_139, x_138); lean_dec(x_138); x_141 = l_Lean_Syntax_mkNameLit(x_140, x_121); x_142 = lean_array_push(x_123, x_141); -x_143 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_143 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_144 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_144, 0, x_121); lean_ctor_set(x_144, 1, x_143); @@ -13424,15 +14206,15 @@ lean_inc(x_155); x_157 = l_Lean_addMacroScope(x_155, x_156, x_96); x_158 = lean_box(0); x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; -x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; +x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; lean_inc(x_94); lean_ctor_set(x_1, 3, x_160); lean_ctor_set(x_1, 2, x_157); lean_ctor_set(x_1, 1, x_159); lean_ctor_set(x_1, 0, x_94); -x_161 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_161 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_162 = l_Lean_addMacroScope(x_155, x_161, x_96); -x_163 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_163 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_94); x_164 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_164, 0, x_94); @@ -13441,7 +14223,7 @@ lean_ctor_set(x_164, 2, x_162); lean_ctor_set(x_164, 3, x_158); lean_inc(x_87); x_165 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_158, x_87); -x_166 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_88); +x_166 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(x_88); x_167 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_168 = lean_array_push(x_167, x_164); x_169 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; @@ -13495,14 +14277,14 @@ lean_dec(x_87); x_191 = lean_ctor_get(x_165, 0); lean_inc(x_191); lean_dec(x_165); -x_192 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_192 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_193 = l_String_intercalate(x_192, x_191); -x_194 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_194 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_195 = lean_string_append(x_194, x_193); lean_dec(x_193); x_196 = l_Lean_Syntax_mkNameLit(x_195, x_175); x_197 = lean_array_push(x_177, x_196); -x_198 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_198 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_175); lean_ctor_set(x_199, 1, x_198); @@ -13599,16 +14381,16 @@ lean_inc(x_225); x_227 = l_Lean_addMacroScope(x_225, x_226, x_219); x_228 = lean_box(0); x_229 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__3; -x_230 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__8; +x_230 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; lean_inc(x_217); x_231 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_231, 0, x_217); lean_ctor_set(x_231, 1, x_229); lean_ctor_set(x_231, 2, x_227); lean_ctor_set(x_231, 3, x_230); -x_232 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_232 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; x_233 = l_Lean_addMacroScope(x_225, x_232, x_219); -x_234 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_234 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_217); x_235 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_235, 0, x_217); @@ -13617,7 +14399,7 @@ lean_ctor_set(x_235, 2, x_233); lean_ctor_set(x_235, 3, x_228); lean_inc(x_211); x_236 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_228, x_211); -x_237 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__28(x_212); +x_237 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__30(x_212); x_238 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_239 = lean_array_push(x_238, x_235); x_240 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; @@ -13675,14 +14457,14 @@ lean_dec(x_211); x_262 = lean_ctor_get(x_236, 0); lean_inc(x_262); lean_dec(x_236); -x_263 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_263 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_264 = l_String_intercalate(x_263, x_262); -x_265 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_265 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_266 = lean_string_append(x_265, x_264); lean_dec(x_264); x_267 = l_Lean_Syntax_mkNameLit(x_266, x_246); x_268 = lean_array_push(x_248, x_267); -x_269 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_269 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_270 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_270, 0, x_246); lean_ctor_set(x_270, 1, x_269); @@ -14102,13 +14884,13 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___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_2; -x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_7; +x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__29___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_2); +return x_7; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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) { @@ -16965,7 +17747,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -16975,7 +17757,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__50 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -17121,11 +17903,11 @@ lean_inc(x_18); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_18); lean_ctor_set(x_41, 1, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; lean_inc(x_20); lean_inc(x_26); x_43 = l_Lean_addMacroScope(x_26, x_42, x_20); -x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; lean_inc(x_18); x_45 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_45, 0, x_18); @@ -17290,15 +18072,15 @@ lean_dec(x_11); x_180 = lean_ctor_get(x_93, 0); lean_inc(x_180); lean_dec(x_93); -x_181 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_181 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_182 = l_String_intercalate(x_181, x_180); -x_183 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_183 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_184 = lean_string_append(x_183, x_182); lean_dec(x_182); x_185 = lean_box(2); x_186 = l_Lean_Syntax_mkNameLit(x_184, x_185); x_187 = lean_array_push(x_46, x_186); -x_188 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_188 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_189 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_189, 0, x_185); lean_ctor_set(x_189, 1, x_188); @@ -18844,7 +19626,7 @@ return x_239; } } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1() { _start: { lean_object* x_1; @@ -18852,16 +19634,16 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_stxQuot_expand), 8, return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -18871,17 +19653,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -18891,7 +19673,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4() { _start: { lean_object* x_1; @@ -18899,47 +19681,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9() { _start: { lean_object* x_1; @@ -18947,27 +19729,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5447u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5655u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12() { _start: { lean_object* x_1; @@ -18975,31 +19757,31 @@ x_1 = l_Lean_Elab_Term_termElabAttribute; return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19007,11 +19789,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19019,13 +19801,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2; x_4 = lean_unsigned_to_nat(31u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19035,37 +19817,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19075,41 +19857,41 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5453u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5661u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19117,11 +19899,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19129,13 +19911,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2; x_4 = lean_unsigned_to_nat(30u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19145,37 +19927,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1() { _start: { lean_object* x_1; @@ -19183,61 +19965,61 @@ x_1 = lean_mk_string_from_bytes("funBinder", 9); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5459u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5667u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(247u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19245,11 +20027,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(247u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19257,13 +20039,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2; x_4 = lean_unsigned_to_nat(40u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19273,37 +20055,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1() { _start: { lean_object* x_1; @@ -19311,61 +20093,61 @@ x_1 = lean_mk_string_from_bytes("bracketedBinder", 15); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5465u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5673u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(254u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19373,11 +20155,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(254u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19385,13 +20167,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2; x_4 = lean_unsigned_to_nat(46u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19401,81 +20183,81 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5471u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5679u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19483,11 +20265,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19495,13 +20277,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2; x_4 = lean_unsigned_to_nat(41u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19511,37 +20293,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19551,51 +20333,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5477u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5685u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19603,11 +20385,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19615,13 +20397,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2; x_4 = lean_unsigned_to_nat(32u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19631,81 +20413,81 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5483u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5691u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19713,11 +20495,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19725,13 +20507,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19741,37 +20523,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19781,51 +20563,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5489u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5697u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(252u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19833,11 +20615,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(252u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19845,13 +20627,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2; x_4 = lean_unsigned_to_nat(34u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19861,37 +20643,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19901,51 +20683,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5495u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5703u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(259u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19953,11 +20735,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(259u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19965,13 +20747,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -19981,37 +20763,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20021,51 +20803,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5501u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5709u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(254u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20073,11 +20855,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(254u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20085,13 +20867,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -20101,37 +20883,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20141,51 +20923,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5507u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5715u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(255u); +x_1 = lean_unsigned_to_nat(261u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20193,11 +20975,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(255u); +x_1 = lean_unsigned_to_nat(261u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20205,13 +20987,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -20221,37 +21003,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20261,51 +21043,51 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5513u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5721u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(256u); +x_1 = lean_unsigned_to_nat(262u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20313,11 +21095,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(256u); +x_1 = lean_unsigned_to_nat(262u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20325,13 +21107,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2; x_4 = lean_unsigned_to_nat(37u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -20341,37 +21123,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20381,41 +21163,41 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5519u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5727u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(257u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20423,11 +21205,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(257u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20435,13 +21217,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2; x_4 = lean_unsigned_to_nat(37u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -20451,37 +21233,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20491,41 +21273,41 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10; -x_2 = lean_unsigned_to_nat(5525u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10; +x_2 = lean_unsigned_to_nat(5733u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3; 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(258u); +x_1 = lean_unsigned_to_nat(264u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20533,11 +21315,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(258u); +x_1 = lean_unsigned_to_nat(264u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20545,13 +21327,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2; x_4 = lean_unsigned_to_nat(33u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -20561,23 +21343,23 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___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; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } @@ -21030,7 +21812,7 @@ x_9 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_T x_10 = lean_array_push(x_9, x_6); x_11 = lean_array_push(x_10, x_1); x_12 = lean_array_push(x_11, x_8); -x_13 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1; +x_13 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1; x_14 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_14, 0, x_2); lean_ctor_set(x_14, 1, x_13); @@ -21115,7 +21897,7 @@ static lean_object* _init_l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -21125,7 +21907,7 @@ static lean_object* _init_l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -21797,7 +22579,7 @@ x_43 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotatio lean_inc(x_34); lean_inc(x_39); x_44 = l_Lean_addMacroScope(x_39, x_43, x_34); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +x_45 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4; lean_inc(x_3); x_46 = l_Lean_Name_str___override(x_3, x_45); x_47 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; @@ -21983,7 +22765,7 @@ x_108 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotati lean_inc(x_99); lean_inc(x_104); x_109 = l_Lean_addMacroScope(x_104, x_108, x_99); -x_110 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +x_110 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4; lean_inc(x_3); x_111 = l_Lean_Name_str___override(x_3, x_110); x_112 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; @@ -23679,7 +24461,7 @@ static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_Quotation_0__Le _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -23776,7 +24558,7 @@ x_45 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotat lean_inc(x_23); lean_inc(x_28); x_46 = l_Lean_addMacroScope(x_28, x_45, x_23); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +x_47 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4; lean_inc(x_2); x_48 = l_Lean_Name_str___override(x_2, x_47); x_49 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__6; @@ -23887,12 +24669,12 @@ lean_dec(x_18); x_89 = lean_ctor_get(x_60, 0); lean_inc(x_89); lean_dec(x_60); -x_90 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; +x_90 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15; lean_inc(x_3); x_91 = l_Lean_Name_str___override(x_3, x_90); -x_92 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_92 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_93 = l_String_intercalate(x_92, x_89); -x_94 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_94 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_95 = lean_string_append(x_94, x_93); lean_dec(x_93); x_96 = lean_box(2); @@ -24021,7 +24803,7 @@ x_144 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota lean_inc(x_121); lean_inc(x_126); x_145 = l_Lean_addMacroScope(x_126, x_144, x_121); -x_146 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +x_146 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4; lean_inc(x_2); x_147 = l_Lean_Name_str___override(x_2, x_146); x_148 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__6; @@ -24132,12 +24914,12 @@ lean_dec(x_116); x_188 = lean_ctor_get(x_159, 0); lean_inc(x_188); lean_dec(x_159); -x_189 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; +x_189 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15; lean_inc(x_3); x_190 = l_Lean_Name_str___override(x_3, x_189); -x_191 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_191 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_192 = l_String_intercalate(x_191, x_188); -x_193 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_193 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_194 = lean_string_append(x_193, x_192); lean_dec(x_192); x_195 = lean_box(2); @@ -25404,7 +26186,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25453,7 +26235,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25463,7 +26245,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25512,7 +26294,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25522,7 +26304,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25571,7 +26353,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25581,7 +26363,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -25740,16 +26522,16 @@ lean_dec(x_5); x_55 = lean_ctor_get(x_43, 0); lean_inc(x_55); lean_dec(x_43); -x_56 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_56 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_57 = l_String_intercalate(x_56, x_55); -x_58 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_58 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_59 = lean_string_append(x_58, x_57); lean_dec(x_57); x_60 = lean_box(2); x_61 = l_Lean_Syntax_mkNameLit(x_59, x_60); x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_63 = lean_array_push(x_62, x_61); -x_64 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_64 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_60); lean_ctor_set(x_65, 1, x_64); @@ -25862,16 +26644,16 @@ lean_dec(x_93); x_106 = lean_ctor_get(x_94, 0); lean_inc(x_106); lean_dec(x_94); -x_107 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_107 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_108 = l_String_intercalate(x_107, x_106); -x_109 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_109 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_110 = lean_string_append(x_109, x_108); lean_dec(x_108); x_111 = lean_box(2); x_112 = l_Lean_Syntax_mkNameLit(x_110, x_111); x_113 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_114 = lean_array_push(x_113, x_112); -x_115 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_115 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_116 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_116, 0, x_111); lean_ctor_set(x_116, 1, x_115); @@ -26064,16 +26846,16 @@ lean_dec(x_5); x_190 = lean_ctor_get(x_178, 0); lean_inc(x_190); lean_dec(x_178); -x_191 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_191 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_192 = l_String_intercalate(x_191, x_190); -x_193 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_193 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_194 = lean_string_append(x_193, x_192); lean_dec(x_192); x_195 = lean_box(2); x_196 = l_Lean_Syntax_mkNameLit(x_194, x_195); x_197 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_198 = lean_array_push(x_197, x_196); -x_199 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_199 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_200 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_200, 0, x_195); lean_ctor_set(x_200, 1, x_199); @@ -26185,16 +26967,16 @@ lean_dec(x_5); x_240 = lean_ctor_get(x_228, 0); lean_inc(x_240); lean_dec(x_228); -x_241 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_241 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_242 = l_String_intercalate(x_241, x_240); -x_243 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_243 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_244 = lean_string_append(x_243, x_242); lean_dec(x_242); x_245 = lean_box(2); x_246 = l_Lean_Syntax_mkNameLit(x_244, x_245); x_247 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_248 = lean_array_push(x_247, x_246); -x_249 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_249 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_250 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_250, 0, x_245); lean_ctor_set(x_250, 1, x_249); @@ -26296,16 +27078,16 @@ lean_dec(x_5); x_311 = lean_ctor_get(x_278, 0); lean_inc(x_311); lean_dec(x_278); -x_312 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_312 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; x_313 = l_String_intercalate(x_312, x_311); -x_314 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_314 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18; x_315 = lean_string_append(x_314, x_313); lean_dec(x_313); x_316 = lean_box(2); x_317 = l_Lean_Syntax_mkNameLit(x_315, x_316); x_318 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_319 = lean_array_push(x_318, x_317); -x_320 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_320 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16; x_321 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_321, 0, x_316); lean_ctor_set(x_321, 1, x_320); @@ -27547,7 +28329,7 @@ lean_ctor_set(x_39, 1, x_38); x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; lean_inc(x_1); x_41 = l_Lean_Name_str___override(x_1, x_40); -x_42 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_42 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; lean_inc(x_21); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_21); @@ -27733,9 +28515,9 @@ x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_21); lean_ctor_set(x_112, 1, x_2); lean_ctor_set(x_112, 2, x_111); -x_113 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_113 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; x_114 = l_Lean_addMacroScope(x_28, x_113, x_23); -x_115 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +x_115 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; lean_inc(x_3); x_116 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_116, 0, x_115); @@ -27743,7 +28525,7 @@ lean_ctor_set(x_116, 1, x_3); x_117 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_3); -x_118 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; +x_118 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; lean_inc(x_21); x_119 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_119, 0, x_21); @@ -27983,7 +28765,7 @@ lean_ctor_set(x_200, 1, x_199); x_201 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; lean_inc(x_1); x_202 = l_Lean_Name_str___override(x_1, x_201); -x_203 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_203 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17; lean_inc(x_21); x_204 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_204, 0, x_21); @@ -28169,9 +28951,9 @@ x_273 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_273, 0, x_21); lean_ctor_set(x_273, 1, x_2); lean_ctor_set(x_273, 2, x_272); -x_274 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_274 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; x_275 = l_Lean_addMacroScope(x_189, x_274, x_23); -x_276 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +x_276 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; lean_inc(x_3); x_277 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_277, 0, x_276); @@ -28179,7 +28961,7 @@ lean_ctor_set(x_277, 1, x_3); x_278 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_278, 0, x_277); lean_ctor_set(x_278, 1, x_3); -x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; +x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; lean_inc(x_21); x_280 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_280, 0, x_21); @@ -29000,7 +29782,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -29010,7 +29792,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -29731,10 +30513,10 @@ x_208 = lean_ctor_get(x_206, 0); lean_inc(x_208); lean_dec(x_206); x_209 = lean_environment_main_module(x_208); -x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; lean_inc(x_204); x_211 = l_Lean_addMacroScope(x_209, x_210, x_204); -x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; lean_inc(x_2); x_213 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_213, 0, x_212); @@ -29743,7 +30525,7 @@ lean_inc(x_2); x_214 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_214, 0, x_213); lean_ctor_set(x_214, 1, x_2); -x_215 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; +x_215 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; x_216 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_216, 0, x_202); lean_ctor_set(x_216, 1, x_215); @@ -29939,7 +30721,7 @@ lean_ctor_set(x_297, 3, x_2); lean_inc(x_244); x_298 = lean_array_push(x_285, x_244); x_299 = lean_array_push(x_298, x_297); -x_300 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_300 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_inc(x_220); x_301 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_301, 0, x_220); @@ -29975,7 +30757,7 @@ x_312 = lean_array_push(x_272, x_309); x_313 = lean_array_push(x_312, x_311); lean_inc(x_271); x_314 = lean_array_push(x_313, x_271); -x_315 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1; +x_315 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1; lean_inc(x_220); x_316 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_316, 0, x_220); @@ -30289,7 +31071,7 @@ lean_ctor_set(x_451, 3, x_2); lean_inc(x_398); x_452 = lean_array_push(x_439, x_398); x_453 = lean_array_push(x_452, x_451); -x_454 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_454 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_inc(x_220); x_455 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_455, 0, x_220); @@ -30325,7 +31107,7 @@ x_466 = lean_array_push(x_426, x_463); x_467 = lean_array_push(x_466, x_465); lean_inc(x_425); x_468 = lean_array_push(x_467, x_425); -x_469 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1; +x_469 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1; lean_inc(x_220); x_470 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_470, 0, x_220); @@ -31045,7 +31827,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -31086,7 +31868,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -31096,7 +31878,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -31145,7 +31927,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -31155,7 +31937,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; @@ -32089,7 +32871,7 @@ lean_inc(x_389); x_403 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_403, 0, x_389); lean_ctor_set(x_403, 1, x_402); -x_404 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +x_404 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; lean_inc(x_391); lean_inc(x_396); x_405 = l_Lean_addMacroScope(x_396, x_404, x_391); @@ -32101,7 +32883,7 @@ lean_inc(x_3); x_407 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_407, 0, x_406); lean_ctor_set(x_407, 1, x_3); -x_408 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; +x_408 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; lean_inc(x_389); x_409 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_409, 0, x_389); @@ -32326,7 +33108,7 @@ lean_inc(x_389); x_503 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_503, 0, x_389); lean_ctor_set(x_503, 1, x_502); -x_504 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +x_504 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; lean_inc(x_391); lean_inc(x_496); x_505 = l_Lean_addMacroScope(x_496, x_504, x_391); @@ -32338,7 +33120,7 @@ lean_inc(x_3); x_507 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_507, 0, x_506); lean_ctor_set(x_507, 1, x_3); -x_508 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; +x_508 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; lean_inc(x_389); x_509 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_509, 0, x_389); @@ -33103,7 +33885,7 @@ x_41 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; lean_inc(x_17); lean_inc(x_22); x_42 = l_Lean_addMacroScope(x_22, x_41, x_17); -x_43 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; +x_43 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; x_44 = l_Lean_Name_str___override(x_2, x_43); x_45 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_46 = l_Lean_Name_str___override(x_44, x_45); @@ -33241,7 +34023,7 @@ x_106 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; lean_inc(x_17); lean_inc(x_87); x_107 = l_Lean_addMacroScope(x_87, x_106, x_17); -x_108 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; +x_108 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; x_109 = l_Lean_Name_str___override(x_2, x_108); x_110 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_111 = l_Lean_Name_str___override(x_109, x_110); @@ -33573,7 +34355,7 @@ lean_object* x_88; lean_object* x_89; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_89 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_1, x_88, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_89; } @@ -33584,7 +34366,7 @@ lean_object* x_90; lean_object* x_91; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_91 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_1, x_90, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_91; } @@ -33767,7 +34549,7 @@ lean_object* x_136; lean_object* x_137; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_136 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_136 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_137 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_1, x_136, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_137; } @@ -33778,7 +34560,7 @@ lean_object* x_138; lean_object* x_139; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_138 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; +x_138 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; x_139 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_1, x_138, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_139; } @@ -39580,7 +40362,7 @@ lean_inc(x_41); x_47 = lean_array_push(x_46, x_41); lean_inc(x_47); x_48 = lean_array_push(x_47, x_45); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_inc(x_29); x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_29); @@ -39758,7 +40540,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4; -x_3 = lean_unsigned_to_nat(572u); +x_3 = lean_unsigned_to_nat(578u); x_4 = lean_unsigned_to_nat(12u); x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -44524,7 +45306,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__3; @@ -44554,7 +45336,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(652u); +x_1 = lean_unsigned_to_nat(658u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44566,7 +45348,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(653u); +x_1 = lean_unsigned_to_nat(659u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44594,7 +45376,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(652u); +x_1 = lean_unsigned_to_nat(658u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44606,7 +45388,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(652u); +x_1 = lean_unsigned_to_nat(658u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44748,7 +45530,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUn _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__5___closed__1; x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused___closed__3; @@ -44760,7 +45542,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfU _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(655u); +x_1 = lean_unsigned_to_nat(661u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44772,7 +45554,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfU _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(658u); +x_1 = lean_unsigned_to_nat(664u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44800,7 +45582,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfU _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(655u); +x_1 = lean_unsigned_to_nat(661u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44812,7 +45594,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfU _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(655u); +x_1 = lean_unsigned_to_nat(661u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44858,7 +45640,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16119_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16327_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -45118,6 +45900,58 @@ l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__1 = _init_l_Lean_E lean_mark_persistent(l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__1); l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__2 = _init_l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_tryAddSyntaxNodeKindInfo___closed__2); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__1); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__2); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__3); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__4); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__5); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__6); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__7); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__8); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__9); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__10); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__11); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__12); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__13 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__13); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__14); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__15); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__16); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__17); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__1___closed__18); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__1); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__2); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__3); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__4); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__5); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__6); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__7 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__7); +l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8 = _init_l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_instQuotePreresolvedStrAnonymous___lambda__2___closed__8); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__2(); @@ -45154,14 +45988,6 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2(); @@ -45394,10 +46220,6 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Qu lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2(); @@ -45476,8 +46298,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__39); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__40); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__2(); @@ -45534,6 +46354,10 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__27); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3); l_Lean_Elab_Term_Quotation_addNamedQuotInfo___closed__1 = _init_l_Lean_Elab_Term_Quotation_addNamedQuotInfo___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_addNamedQuotInfo___closed__1); l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__6___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__6___closed__1(); @@ -45868,326 +46692,326 @@ l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__E lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61); l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62(); lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62); -l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1 = _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__5); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__6); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__7); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__8); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__9); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__10); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__11); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__12); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447____closed__13); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447_(lean_io_mk_world()); +l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1 = _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__5); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__6); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__7); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__8); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__9); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__10); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__11); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__12); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655____closed__13); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5447__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5655__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5453__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5661__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459____closed__5); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667____closed__5); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5459__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5667__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465____closed__5); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673____closed__5); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5465__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5673__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5471__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5679__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5477__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5685__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5483__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5691__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5489__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5697__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5495__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5703__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5501__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5709__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5507__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5715__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5513__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5721__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5519__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5727__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5525__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_5733__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1(); @@ -46673,7 +47497,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused res = l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16119_(lean_io_mk_world()); +res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_16327_(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/InferType.c b/stage0/stdlib/Lean/Meta/InferType.c index 282ad0db5c78..8bf685cd0aad 100644 --- a/stage0/stdlib/Lean/Meta/InferType.c +++ b/stage0/stdlib/Lean/Meta/InferType.c @@ -51,6 +51,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_throwFunctionExpected(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_inferTypeImp_infer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at___private_Lean_Meta_InferType_0__Lean_Meta_checkInferTypeCache___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_isTypeFormerType_go___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwTypeExcepted___rarg___closed__1; static lean_object* l_Lean_Expr_instantiateBetaRevRange_visit___closed__12; @@ -69,6 +70,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Expr_0__Lean_Expr_withAppRevAux___at_Lean_Expr_instantiateBetaRevRange_visit___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); +LEAN_EXPORT uint8_t l_Lean_Meta_isTypeFormerTypeQuick(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__1___boxed(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___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -142,6 +144,7 @@ static lean_object* l_Lean_Meta_throwUnknownMVar___rarg___closed__2; static lean_object* l_Lean_Meta_inferTypeImp_infer___closed__2; uint64_t lean_uint64_of_nat(lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_isTypeFormerType_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_inferFVarType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; size_t lean_usize_modn(size_t, lean_object*); @@ -186,6 +189,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_InferType_0_ uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_inferLambdaType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isTypeFormerTypeQuick___boxed(lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); lean_object* l_Lean_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); lean_object* l_Lean_mkAppRev(lean_object*, lean_object*); @@ -210,6 +214,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_isPropQuic static lean_object* l_Lean_Meta_throwFunctionExpected___rarg___closed__4; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwUnknownMVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_StateT_instMonadStateT___rarg(lean_object*); @@ -5031,450 +5036,6 @@ x_8 = l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_InferType_0__Lean_ return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg(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: -{ -lean_object* x_10; 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_20 = lean_st_ref_get(x_8, x_9); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_st_ref_get(x_6, x_21); -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(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(x_8, x_24); -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_inc(x_27); -x_29 = l_Lean_Expr_fvar___override(x_27); -x_30 = !lean_is_exclusive(x_5); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_5, 1); -x_32 = lean_local_ctx_mk_local_decl(x_31, x_27, x_1, x_3, x_2); -lean_ctor_set(x_5, 1, x_32); -lean_inc(x_8); -lean_inc(x_6); -x_33 = lean_apply_6(x_4, x_29, x_5, x_6, x_7, x_8, x_28); -if (lean_obj_tag(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; uint8_t x_41; -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); -x_36 = lean_st_ref_get(x_8, x_35); -lean_dec(x_8); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_st_ref_take(x_6, x_37); -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_is_exclusive(x_39); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_39, 1); -lean_dec(x_42); -lean_ctor_set(x_39, 1, x_25); -x_43 = lean_st_ref_set(x_6, x_39, x_40); -lean_dec(x_6); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) -{ -lean_object* x_45; -x_45 = lean_ctor_get(x_43, 0); -lean_dec(x_45); -lean_ctor_set(x_43, 0, x_34); -x_10 = x_43; -goto block_19; -} -else -{ -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_34); -lean_ctor_set(x_47, 1, x_46); -x_10 = x_47; -goto block_19; -} -} -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_ctor_get(x_39, 0); -x_49 = lean_ctor_get(x_39, 2); -x_50 = lean_ctor_get(x_39, 3); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_39); -x_51 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_51, 0, x_48); -lean_ctor_set(x_51, 1, x_25); -lean_ctor_set(x_51, 2, x_49); -lean_ctor_set(x_51, 3, x_50); -x_52 = lean_st_ref_set(x_6, x_51, x_40); -lean_dec(x_6); -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_34); -lean_ctor_set(x_55, 1, x_53); -x_10 = x_55; -goto block_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; uint8_t x_63; -x_56 = lean_ctor_get(x_33, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_33, 1); -lean_inc(x_57); -lean_dec(x_33); -x_58 = lean_st_ref_get(x_8, x_57); -lean_dec(x_8); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_st_ref_take(x_6, x_59); -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = !lean_is_exclusive(x_61); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_64 = lean_ctor_get(x_61, 1); -lean_dec(x_64); -lean_ctor_set(x_61, 1, x_25); -x_65 = lean_st_ref_set(x_6, x_61, x_62); -lean_dec(x_6); -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_tag(x_65, 1); -lean_ctor_set(x_65, 0, x_56); -x_10 = x_65; -goto block_19; -} -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); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_56); -lean_ctor_set(x_69, 1, x_68); -x_10 = x_69; -goto block_19; -} -} -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; -x_70 = lean_ctor_get(x_61, 0); -x_71 = lean_ctor_get(x_61, 2); -x_72 = lean_ctor_get(x_61, 3); -lean_inc(x_72); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_61); -x_73 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_25); -lean_ctor_set(x_73, 2, x_71); -lean_ctor_set(x_73, 3, x_72); -x_74 = lean_st_ref_set(x_6, x_73, x_62); -lean_dec(x_6); -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); -} -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); -} else { - x_77 = x_76; - lean_ctor_set_tag(x_77, 1); -} -lean_ctor_set(x_77, 0, x_56); -lean_ctor_set(x_77, 1, x_75); -x_10 = x_77; -goto block_19; -} -} -} -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; -x_78 = lean_ctor_get(x_5, 0); -x_79 = lean_ctor_get(x_5, 1); -x_80 = lean_ctor_get(x_5, 2); -x_81 = lean_ctor_get(x_5, 3); -x_82 = lean_ctor_get(x_5, 4); -x_83 = lean_ctor_get(x_5, 5); -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_dec(x_5); -x_84 = lean_local_ctx_mk_local_decl(x_79, x_27, x_1, x_3, x_2); -x_85 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_85, 0, x_78); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_80); -lean_ctor_set(x_85, 3, x_81); -lean_ctor_set(x_85, 4, x_82); -lean_ctor_set(x_85, 5, x_83); -lean_inc(x_8); -lean_inc(x_6); -x_86 = lean_apply_6(x_4, x_29, x_85, x_6, x_7, x_8, x_28); -if (lean_obj_tag(x_86) == 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; 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; -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_st_ref_get(x_8, x_88); -lean_dec(x_8); -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -lean_dec(x_89); -x_91 = lean_st_ref_take(x_6, x_90); -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_94 = lean_ctor_get(x_92, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_92, 2); -lean_inc(x_95); -x_96 = lean_ctor_get(x_92, 3); -lean_inc(x_96); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - lean_ctor_release(x_92, 2); - lean_ctor_release(x_92, 3); - x_97 = x_92; -} else { - lean_dec_ref(x_92); - x_97 = lean_box(0); -} -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(0, 4, 0); -} else { - x_98 = x_97; -} -lean_ctor_set(x_98, 0, x_94); -lean_ctor_set(x_98, 1, x_25); -lean_ctor_set(x_98, 2, x_95); -lean_ctor_set(x_98, 3, x_96); -x_99 = lean_st_ref_set(x_6, x_98, x_93); -lean_dec(x_6); -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); -} -if (lean_is_scalar(x_101)) { - x_102 = lean_alloc_ctor(0, 2, 0); -} else { - x_102 = x_101; -} -lean_ctor_set(x_102, 0, x_87); -lean_ctor_set(x_102, 1, x_100); -x_10 = x_102; -goto block_19; -} -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; -x_103 = lean_ctor_get(x_86, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_86, 1); -lean_inc(x_104); -lean_dec(x_86); -x_105 = lean_st_ref_get(x_8, x_104); -lean_dec(x_8); -x_106 = lean_ctor_get(x_105, 1); -lean_inc(x_106); -lean_dec(x_105); -x_107 = lean_st_ref_take(x_6, x_106); -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_ctor_get(x_108, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_108, 2); -lean_inc(x_111); -x_112 = lean_ctor_get(x_108, 3); -lean_inc(x_112); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - lean_ctor_release(x_108, 2); - lean_ctor_release(x_108, 3); - x_113 = x_108; -} else { - lean_dec_ref(x_108); - x_113 = lean_box(0); -} -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(0, 4, 0); -} else { - x_114 = x_113; -} -lean_ctor_set(x_114, 0, x_110); -lean_ctor_set(x_114, 1, x_25); -lean_ctor_set(x_114, 2, x_111); -lean_ctor_set(x_114, 3, x_112); -x_115 = lean_st_ref_set(x_6, x_114, x_109); -lean_dec(x_6); -x_116 = lean_ctor_get(x_115, 1); -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); -} -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(1, 2, 0); -} else { - x_118 = x_117; - lean_ctor_set_tag(x_118, 1); -} -lean_ctor_set(x_118, 0, x_103); -lean_ctor_set(x_118, 1, x_116); -x_10 = x_118; -goto block_19; -} -} -block_19: -{ -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) -{ -return x_10; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -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_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg___boxed), 9, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___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: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg(x_1, 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_throwError___at_Lean_Meta_throwUnknownMVar___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) { _start: { @@ -13569,337 +13130,583 @@ return x_59; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_isTypeFormerType(lean_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___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg(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: { -lean_object* x_7; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_7 = l_Lean_Meta_whnfD(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(x_8, x_9); +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_inc(x_11); +x_13 = l_Lean_Expr_fvar___override(x_11); +x_14 = !lean_is_exclusive(x_5); +if (x_14 == 0) { -lean_object* x_8; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -switch (lean_obj_tag(x_8)) { -case 3: +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_local_ctx_mk_local_decl(x_15, x_11, x_1, x_3, x_2); +lean_ctor_set(x_5, 1, x_16); +x_17 = lean_apply_6(x_4, x_13, x_5, x_6, x_7, x_8, x_12); +return x_17; +} +else { -uint8_t x_9; -lean_dec(x_8); +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); +x_20 = lean_ctor_get(x_5, 2); +x_21 = lean_ctor_get(x_5, 3); +x_22 = lean_ctor_get(x_5, 4); +x_23 = lean_ctor_get(x_5, 5); +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_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +x_24 = lean_local_ctx_mk_local_decl(x_19, x_11, x_1, x_3, x_2); +x_25 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_21); +lean_ctor_set(x_25, 4, x_22); +lean_ctor_set(x_25, 5, x_23); +x_26 = lean_apply_6(x_4, x_13, x_25, x_6, x_7, x_8, x_12); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg___boxed), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___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: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); lean_dec(x_2); -x_9 = !lean_is_exclusive(x_7); -if (x_9 == 0) +x_11 = l___private_Lean_Meta_InferType_0__Lean_Meta_withLocalDecl_x27___rarg(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_isTypeFormerTypeQuick(lean_object* x_1) { +_start: { -lean_object* x_10; uint8_t x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 0); -lean_dec(x_10); -x_11 = 1; -x_12 = lean_box(x_11); -lean_ctor_set(x_7, 0, x_12); -return x_7; +switch (lean_obj_tag(x_1)) { +case 3: +{ +uint8_t x_2; +x_2 = 1; +return x_2; } -else +case 7: { -lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -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_13); -return x_16; +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 2); +x_1 = x_3; +goto _start; +} +default: +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} } } +} +LEAN_EXPORT lean_object* l_Lean_Meta_isTypeFormerTypeQuick___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_isTypeFormerTypeQuick(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_isTypeFormerType_go___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_Lean_Meta_isTypeFormerType_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_object* x_7) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 3: +{ +uint8_t x_8; lean_object* x_9; lean_object* 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); +x_8 = 1; +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_7); +return x_10; +} case 7: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; 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_17 = lean_ctor_get(x_7, 1); +lean_object* x_11; 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; lean_object* x_27; lean_object* x_28; +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_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +x_15 = lean_expr_instantiate_rev(x_12, x_2); +lean_dec(x_12); +x_16 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(x_6, x_7); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_7); -x_18 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_8, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_8, 2); +lean_dec(x_16); +lean_inc(x_17); +x_19 = l_Lean_Expr_fvar___override(x_17); +x_20 = lean_ctor_get(x_3, 0); lean_inc(x_20); -x_21 = lean_ctor_get_uint8(x_8, sizeof(void*)*3 + 8); -lean_dec(x_8); -x_32 = lean_st_ref_get(x_5, x_17); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = lean_st_ref_get(x_3, x_33); -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, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(x_5, x_36); -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); -lean_inc(x_39); -x_41 = l_Lean_Expr_fvar___override(x_39); -x_42 = lean_ctor_get(x_2, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_2, 1); -lean_inc(x_43); -x_44 = lean_local_ctx_mk_local_decl(x_43, x_39, x_18, x_19, x_21); -x_45 = lean_ctor_get(x_2, 2); -lean_inc(x_45); -x_46 = lean_ctor_get(x_2, 3); -lean_inc(x_46); -x_47 = lean_ctor_get(x_2, 4); -lean_inc(x_47); -x_48 = lean_ctor_get(x_2, 5); -lean_inc(x_48); +x_21 = lean_ctor_get(x_3, 1); +lean_inc(x_21); +x_22 = lean_local_ctx_mk_local_decl(x_21, x_17, x_11, x_15, x_14); +x_23 = lean_ctor_get(x_3, 2); +lean_inc(x_23); +x_24 = lean_ctor_get(x_3, 3); +lean_inc(x_24); +x_25 = lean_ctor_get(x_3, 4); +lean_inc(x_25); +x_26 = lean_ctor_get(x_3, 5); +lean_inc(x_26); +lean_dec(x_3); +x_27 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 3, x_24); +lean_ctor_set(x_27, 4, x_25); +lean_ctor_set(x_27, 5, x_26); +x_28 = lean_array_push(x_2, x_19); +x_1 = x_13; +x_2 = x_28; +x_3 = x_27; +x_7 = x_18; +goto _start; +} +default: +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_expr_instantiate_rev(x_1, x_2); lean_dec(x_2); -x_49 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_49, 0, x_42); -lean_ctor_set(x_49, 1, x_44); -lean_ctor_set(x_49, 2, x_45); -lean_ctor_set(x_49, 3, x_46); -lean_ctor_set(x_49, 4, x_47); -lean_ctor_set(x_49, 5, x_48); -x_50 = lean_expr_instantiate1(x_20, x_41); -lean_dec(x_41); -lean_dec(x_20); +lean_dec(x_1); +lean_inc(x_6); lean_inc(x_5); +lean_inc(x_4); lean_inc(x_3); -x_51 = l_Lean_Meta_isTypeFormerType(x_50, x_49, x_3, x_4, x_5, x_40); -if (lean_obj_tag(x_51) == 0) +x_31 = l_Lean_Meta_whnfD(x_30, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_31) == 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; lean_object* x_63; uint8_t x_64; -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_st_ref_get(x_5, x_53); +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +switch (lean_obj_tag(x_32)) { +case 3: +{ +uint8_t x_33; +lean_dec(x_32); +lean_dec(x_6); lean_dec(x_5); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_st_ref_take(x_3, 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); -lean_dec(x_56); -x_59 = lean_ctor_get(x_57, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_57, 2); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 3); -lean_inc(x_61); -lean_dec(x_57); -x_62 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_62, 0, x_59); -lean_ctor_set(x_62, 1, x_37); -lean_ctor_set(x_62, 2, x_60); -lean_ctor_set(x_62, 3, x_61); -x_63 = lean_st_ref_set(x_3, x_62, x_58); +lean_dec(x_4); lean_dec(x_3); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) { -lean_object* x_65; -x_65 = lean_ctor_get(x_63, 0); -lean_dec(x_65); -lean_ctor_set(x_63, 0, x_52); -x_22 = x_63; -goto block_31; +lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_31, 0); +lean_dec(x_34); +x_35 = 1; +x_36 = lean_box(x_35); +lean_ctor_set(x_31, 0, x_36); +return x_31; } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_52); -lean_ctor_set(x_67, 1, x_66); -x_22 = x_67; -goto block_31; +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_31, 1); +lean_inc(x_37); +lean_dec(x_31); +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_37); +return x_40; } } +case 7: +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_31, 1); +lean_inc(x_41); +lean_dec(x_31); +x_42 = l_Lean_Meta_isTypeFormerType_go___closed__1; +x_1 = x_32; +x_2 = x_42; +x_7 = x_41; +goto _start; +} +default: +{ +uint8_t x_44; +lean_dec(x_32); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_44 = !lean_is_exclusive(x_31); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_31, 0); +lean_dec(x_45); +x_46 = 0; +x_47 = lean_box(x_46); +lean_ctor_set(x_31, 0, x_47); +return x_31; +} 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; uint8_t x_80; -x_68 = lean_ctor_get(x_51, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_51, 1); -lean_inc(x_69); -lean_dec(x_51); -x_70 = lean_st_ref_get(x_5, x_69); +lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_ctor_get(x_31, 1); +lean_inc(x_48); +lean_dec(x_31); +x_49 = 0; +x_50 = lean_box(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_48); +return x_51; +} +} +} +} +else +{ +uint8_t x_52; +lean_dec(x_6); lean_dec(x_5); -x_71 = lean_ctor_get(x_70, 1); -lean_inc(x_71); -lean_dec(x_70); -x_72 = lean_st_ref_take(x_3, x_71); -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 = lean_ctor_get(x_73, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_73, 2); -lean_inc(x_76); -x_77 = lean_ctor_get(x_73, 3); -lean_inc(x_77); -lean_dec(x_73); -x_78 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_37); -lean_ctor_set(x_78, 2, x_76); -lean_ctor_set(x_78, 3, x_77); -x_79 = lean_st_ref_set(x_3, x_78, x_74); -lean_dec(x_3); -x_80 = !lean_is_exclusive(x_79); -if (x_80 == 0) -{ -lean_object* x_81; -x_81 = lean_ctor_get(x_79, 0); -lean_dec(x_81); -lean_ctor_set_tag(x_79, 1); -lean_ctor_set(x_79, 0, x_68); -x_22 = x_79; -goto block_31; +lean_dec(x_4); +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_31); +if (x_52 == 0) +{ +return x_31; } else { -lean_object* x_82; lean_object* x_83; -x_82 = lean_ctor_get(x_79, 1); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_68); -lean_ctor_set(x_83, 1, x_82); -x_22 = x_83; -goto block_31; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_31, 0); +x_54 = lean_ctor_get(x_31, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_31); +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; } } -block_31: +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isTypeFormerType(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_22) == 0) +uint8_t x_7; +x_7 = l_Lean_Meta_isTypeFormerTypeQuick(x_1); +if (x_7 == 0) { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_object* x_8; 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_18 = lean_st_ref_get(x_5, x_6); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_get(x_3, 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_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Meta_isTypeFormerType_go___closed__1; +lean_inc(x_5); +lean_inc(x_3); +x_25 = l_Lean_Meta_isTypeFormerType_go(x_1, x_24, x_2, x_3, x_4, x_5, x_22); +if (lean_obj_tag(x_25) == 0) { -return x_22; +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; +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_st_ref_get(x_5, x_27); +lean_dec(x_5); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_st_ref_take(x_3, 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); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_31, 1); +lean_dec(x_34); +lean_ctor_set(x_31, 1, x_23); +x_35 = lean_st_ref_set(x_3, x_31, x_32); +lean_dec(x_3); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +lean_ctor_set(x_35, 0, x_26); +x_8 = x_35; +goto block_17; } else { -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; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_26); +lean_ctor_set(x_39, 1, x_38); +x_8 = x_39; +goto block_17; } } else { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_22); -if (x_27 == 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_object* x_46; lean_object* x_47; +x_40 = lean_ctor_get(x_31, 0); +x_41 = lean_ctor_get(x_31, 2); +x_42 = lean_ctor_get(x_31, 3); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_31); +x_43 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_23); +lean_ctor_set(x_43, 2, x_41); +lean_ctor_set(x_43, 3, x_42); +x_44 = lean_st_ref_set(x_3, x_43, x_32); +lean_dec(x_3); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; +} else { + lean_dec_ref(x_44); + x_46 = lean_box(0); +} +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_46; +} +lean_ctor_set(x_47, 0, x_26); +lean_ctor_set(x_47, 1, x_45); +x_8 = x_47; +goto block_17; +} +} +else { -return x_22; +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_25, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_25, 1); +lean_inc(x_49); +lean_dec(x_25); +x_50 = lean_st_ref_get(x_5, x_49); +lean_dec(x_5); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = lean_st_ref_take(x_3, x_51); +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_is_exclusive(x_53); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_56 = lean_ctor_get(x_53, 1); +lean_dec(x_56); +lean_ctor_set(x_53, 1, x_23); +x_57 = lean_st_ref_set(x_3, x_53, x_54); +lean_dec(x_3); +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) +{ +lean_object* x_59; +x_59 = lean_ctor_get(x_57, 0); +lean_dec(x_59); +lean_ctor_set_tag(x_57, 1); +lean_ctor_set(x_57, 0, x_48); +x_8 = x_57; +goto block_17; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_22, 0); -x_29 = lean_ctor_get(x_22, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_22); -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_60; lean_object* x_61; +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +lean_dec(x_57); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_48); +lean_ctor_set(x_61, 1, x_60); +x_8 = x_61; +goto block_17; } } +else +{ +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; +x_62 = lean_ctor_get(x_53, 0); +x_63 = lean_ctor_get(x_53, 2); +x_64 = lean_ctor_get(x_53, 3); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_53); +x_65 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_23); +lean_ctor_set(x_65, 2, x_63); +lean_ctor_set(x_65, 3, x_64); +x_66 = lean_st_ref_set(x_3, x_65, x_54); +lean_dec(x_3); +x_67 = lean_ctor_get(x_66, 1); +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); +} +if (lean_is_scalar(x_68)) { + x_69 = lean_alloc_ctor(1, 2, 0); +} else { + x_69 = x_68; + lean_ctor_set_tag(x_69, 1); +} +lean_ctor_set(x_69, 0, x_48); +lean_ctor_set(x_69, 1, x_67); +x_8 = x_69; +goto block_17; } } -default: +block_17: { -uint8_t x_84; +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); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_84 = !lean_is_exclusive(x_7); -if (x_84 == 0) +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 { -lean_object* x_85; uint8_t x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_7, 0); -lean_dec(x_85); -x_86 = 0; -x_87 = lean_box(x_86); -lean_ctor_set(x_7, 0, x_87); -return x_7; +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +return x_8; } else { -lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; -x_88 = lean_ctor_get(x_7, 1); -lean_inc(x_88); -lean_dec(x_7); -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; +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; } } } } else { -uint8_t x_92; +uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_92 = !lean_is_exclusive(x_7); -if (x_92 == 0) -{ -return x_7; -} -else -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_7, 0); -x_94 = lean_ctor_get(x_7, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_7); -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_dec(x_1); +x_70 = 1; +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_6); +return x_72; } } } @@ -14058,6 +13865,8 @@ l_Lean_throwMaxRecDepthAt___at_Lean_Meta_inferTypeImp___spec__1___closed__1 = _i lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Meta_inferTypeImp___spec__1___closed__1); l_Lean_throwMaxRecDepthAt___at_Lean_Meta_inferTypeImp___spec__1___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Meta_inferTypeImp___spec__1___closed__2(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Meta_inferTypeImp___spec__1___closed__2); +l_Lean_Meta_isTypeFormerType_go___closed__1 = _init_l_Lean_Meta_isTypeFormerType_go___closed__1(); +lean_mark_persistent(l_Lean_Meta_isTypeFormerType_go___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Std/Data/HashSet.c b/stage0/stdlib/Std/Data/HashSet.c index bd3e57521816..e46fda02346e 100644 --- a/stage0/stdlib/Std/Data/HashSet.c +++ b/stage0/stdlib/Std/Data/HashSet.c @@ -24,14 +24,19 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_foldM___spec uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Std_HashSet_empty___closed__1; lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_foldM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_foldBucketsM(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSet_forM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_find_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg___lambda__1___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 lean_object* l_Std_HashSetImp_reinsertAux___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_HashSetImp_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_replace___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_fold(lean_object*, lean_object*); @@ -44,13 +49,17 @@ lean_object* l_List_foldlM___rarg(lean_object*, lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l_Std_HashSetImp_contains___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSet_fold___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashSetImp___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Std_HashSetImp_moveEntries___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_fold___spec__2___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_fold___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_foldM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSetImp_fold___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_size___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSetImp_forBucketsM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBuckets___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_toList(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_foldM(lean_object*, lean_object*, lean_object*); @@ -58,6 +67,7 @@ LEAN_EXPORT lean_object* l_Std_HashSet_empty(lean_object*, lean_object*, lean_ob size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBuckets___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSetImp_forM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetBucket_update___rarg(lean_object*, size_t, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_erase(lean_object*); @@ -81,6 +91,7 @@ static lean_object* l_Std_mkHashSetImp___rarg___closed__2; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBucketsM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT size_t l_Std_HashSetImp_mkIdx(lean_object*, lean_object*, size_t); +LEAN_EXPORT lean_object* l_Std_HashSet_instForMHashSet(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_toList___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBucketsM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_isEmpty___rarg___boxed(lean_object*); @@ -92,6 +103,7 @@ LEAN_EXPORT lean_object* l_Std_HashSet_find_x3f___rarg(lean_object*, lean_object size_t lean_usize_modn(size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBucketsM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_toList___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_foldM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_toList___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_toArray(lean_object*, lean_object*, lean_object*); @@ -101,9 +113,11 @@ LEAN_EXPORT lean_object* l_List_foldl___at_Std_HashSetImp_moveEntries___spec__1_ LEAN_EXPORT lean_object* l_Std_HashSetImp_contains(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_insert(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_foldM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Std_HashSet_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_toList___spec__2___rarg(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_foldM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBucketsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -112,12 +126,15 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_toArray___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_fold___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_toArray___rarg(lean_object*); +static lean_object* l_Std_HashSet_instForMHashSet___closed__1; LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSet_toList___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSet_fold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_size(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_numBuckets___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSet_instForMHashSet___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_size___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_fold___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSetImp_forM(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_fold___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_numBuckets___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -127,27 +144,35 @@ LEAN_EXPORT lean_object* l_Std_HashSetImp_foldBuckets(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Std_HashSet_instEmptyCollectionHashSet(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_expand___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSet_toArray___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSet_forM(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_foldM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_moveEntries___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +lean_object* l_List_forM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_erase___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashSet___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_find_x3f(lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_reinsertAux(lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_find_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSetImp_forBucketsM(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSet_toArray___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_toArray___spec__2(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSetImp_fold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldBucketsM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_isEmpty___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSetImp_foldBuckets___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_mkIdx___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_instInhabitedHashSet(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_foldBucketsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_HashSet_forM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Std_HashSetImp_foldBuckets___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_foldM___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_fold___spec__2___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_numBuckets(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_foldBuckets___rarg(lean_object*, lean_object*, lean_object*); @@ -155,6 +180,7 @@ LEAN_EXPORT lean_object* l_Std_mkHashSet___boxed(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_toList___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSet_isEmpty(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetBucket_update(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg___boxed(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_Array_foldlMUnsafe_fold___at_Std_HashSet_toArray___spec__2___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetBucket_update___rarg(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { @@ -826,6 +852,305 @@ lean_dec(x_2); return x_8; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___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_Std_HashSetImp_forBucketsM___spec__1___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_Std_HashSetImp_forBucketsM___spec__1___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_dec(x_6); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_array_uget(x_3, x_4); +lean_inc(x_2); +lean_inc(x_1); +x_10 = l_List_forM___rarg(x_1, lean_box(0), x_9, x_2); +x_11 = lean_box_usize(x_4); +x_12 = lean_box_usize(x_5); +x_13 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_2); +lean_closure_set(x_13, 3, x_3); +lean_closure_set(x_13, 4, x_12); +x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_10, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_apply_2(x_16, lean_box(0), x_6); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg___boxed), 6, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_HashSetImp_forBucketsM___rarg(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_array_get_size(x_2); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, 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_4); +lean_dec(x_3); +lean_dec(x_2); +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +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 +{ +uint8_t x_11; +x_11 = lean_nat_dec_le(x_4, x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_4); +lean_dec(x_3); +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, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_box(0); +x_15 = lean_apply_2(x_13, lean_box(0), x_14); +return x_15; +} +else +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = 0; +x_17 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_18 = lean_box(0); +x_19 = l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___rarg(x_1, x_3, x_2, x_16, x_17, x_18); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_HashSetImp_forBucketsM(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_HashSetImp_forBucketsM___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forBucketsM___spec__1___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_Std_HashSetImp_forBucketsM___spec__1___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_Std_HashSetImp_forBucketsM___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) { +_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_Std_HashSetImp_forBucketsM___spec__1___rarg(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___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_Std_HashSetImp_forM___spec__1___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_Std_HashSetImp_forM___spec__1___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_dec(x_6); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_array_uget(x_3, x_4); +lean_inc(x_2); +lean_inc(x_1); +x_10 = l_List_forM___rarg(x_1, lean_box(0), x_9, x_2); +x_11 = lean_box_usize(x_4); +x_12 = lean_box_usize(x_5); +x_13 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_2); +lean_closure_set(x_13, 3, x_3); +lean_closure_set(x_13, 4, x_12); +x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_10, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_apply_2(x_16, lean_box(0), x_6); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg___boxed), 6, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_HashSetImp_forM___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; uint8_t x_7; +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); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_box(0); +x_11 = lean_apply_2(x_9, lean_box(0), x_10); +return x_11; +} +else +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_5, x_5); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +lean_dec(x_4); +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_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_box(0); +x_16 = lean_apply_2(x_14, lean_box(0), x_15); +return x_16; +} +else +{ +size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = 0; +x_18 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_19 = lean_box(0); +x_20 = l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___rarg(x_1, x_2, x_4, x_17, x_18, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_HashSetImp_forM(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_HashSetImp_forM___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSetImp_forM___spec__1___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_Std_HashSetImp_forM___spec__1___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_Std_HashSetImp_forM___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) { +_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_Std_HashSetImp_forM___spec__1___rarg(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; +} +} LEAN_EXPORT lean_object* l_Std_HashSetImp_find_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -1711,6 +2036,193 @@ lean_dec(x_2); return x_5; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___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_Std_HashSet_forM___spec__1___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_Std_HashSet_forM___spec__1___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_dec(x_6); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_array_uget(x_3, x_4); +lean_inc(x_2); +lean_inc(x_1); +x_10 = l_List_forM___rarg(x_1, lean_box(0), x_9, x_2); +x_11 = lean_box_usize(x_4); +x_12 = lean_box_usize(x_5); +x_13 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_2); +lean_closure_set(x_13, 3, x_3); +lean_closure_set(x_13, 4, x_12); +x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_10, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_apply_2(x_16, lean_box(0), x_6); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg___boxed), 6, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_HashSet_forM___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; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +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) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_box(0); +x_11 = lean_apply_2(x_9, lean_box(0), x_10); +return x_11; +} +else +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_5, x_5); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_ctor_get(x_1, 0); +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 = lean_box(0); +x_16 = lean_apply_2(x_14, lean_box(0), x_15); +return x_16; +} +else +{ +size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = 0; +x_18 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_19 = lean_box(0); +x_20 = l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___rarg(x_1, x_3, x_4, x_17, x_18, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_HashSet_forM(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_HashSet_forM___rarg), 3, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Std_HashSet_forM___spec__1___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_Std_HashSet_forM___spec__1___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_Std_HashSet_forM___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) { +_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_Std_HashSet_forM___spec__1___rarg(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_HashSet_forM___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_HashSet_forM(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_HashSet_instForMHashSet___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_HashSet_forM___rarg), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_HashSet_instForMHashSet(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_HashSet_instForMHashSet___closed__1; +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_HashSet_instForMHashSet___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_HashSet_instForMHashSet(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} LEAN_EXPORT lean_object* l_Std_HashSet_size___rarg(lean_object* x_1) { _start: { @@ -2127,6 +2639,8 @@ l_Std_mkHashSetImp___rarg___closed__2 = _init_l_Std_mkHashSetImp___rarg___closed lean_mark_persistent(l_Std_mkHashSetImp___rarg___closed__2); l_Std_HashSet_empty___closed__1 = _init_l_Std_HashSet_empty___closed__1(); lean_mark_persistent(l_Std_HashSet_empty___closed__1); +l_Std_HashSet_instForMHashSet___closed__1 = _init_l_Std_HashSet_instForMHashSet___closed__1(); +lean_mark_persistent(l_Std_HashSet_instForMHashSet___closed__1); l_Std_HashSet_toArray___rarg___closed__1 = _init_l_Std_HashSet_toArray___rarg___closed__1(); lean_mark_persistent(l_Std_HashSet_toArray___rarg___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/tests/lean/ppNotationCode.lean.expected.out b/tests/lean/ppNotationCode.lean.expected.out index b8b2e16bbc12..a15f33852860 100644 --- a/tests/lean/ppNotationCode.lean.expected.out +++ b/tests/lean/ppNotationCode.lean.expected.out @@ -18,7 +18,7 @@ { raw := Lean.Syntax.node info `Lean.Parser.Term.app #[Lean.Syntax.ident info (String.toSubstring "Nat.add") (Lean.addMacroScope mainModule `Nat.add scp) - [(`Nat.add, [])], + [Lean.Syntax.Preresolved.decl `Nat.add [], Lean.Syntax.Preresolved.namespace `Nat.add], Lean.Syntax.node info `null #[lhs.raw, rhs.raw]] }.raw else let discr := x; diff --git a/tests/lean/run/namespaceHyg.lean b/tests/lean/run/namespaceHyg.lean new file mode 100644 index 000000000000..55ec3afdbfa4 --- /dev/null +++ b/tests/lean/run/namespaceHyg.lean @@ -0,0 +1,3 @@ +macro "classical" : command => `(open Classical) + +classical