Documentation

Lean.ScopedEnvExtension

Instances For
    Instances For
      Equations
      • Lean.ScopedEnvExtension.instInhabitedScopedEntries = { default := { map := default } }
      structure Lean.ScopedEnvExtension.StateStack (α : Type) (β : Type) (σ : Type) :
      Instances For
        Equations
        • Lean.ScopedEnvExtension.instInhabitedStateStack = { default := { stateStack := default, scopedEntries := default, newEntries := default } }
        structure Lean.ScopedEnvExtension.Descr (α : Type) (β : Type) (σ : Type) :
        • name : Lake.Name
        • mkInitial : IO σ
        • ofOLeanEntry : σαLean.ImportM β
        • toOLeanEntry : βα
        • addEntry : σβσ
        • finalizeImport : σσ
        Instances For
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  instance Lean.instInhabitedScopedEnvExtension :
                  {a : Type} → [inst : Inhabited a] → {a_1 a_2 : Type} → Inhabited (Lean.ScopedEnvExtension a a_1 a_2)
                  Equations
                  • Lean.instInhabitedScopedEnvExtension = { default := { descr := default, ext := default } }
                  unsafe def Lean.registerScopedEnvExtensionUnsafe {α : Type} {β : Type} {σ : Type} (descr : Lean.ScopedEnvExtension.Descr α β σ) :
                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[implemented_by Lean.registerScopedEnvExtensionUnsafe]
                    opaque Lean.registerScopedEnvExtension {α : Type} {β : Type} {σ : Type} (descr : Lean.ScopedEnvExtension.Descr α β σ) :
                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        def Lean.ScopedEnvExtension.addEntry {α : Type} {β : Type} {σ : Type} (ext : Lean.ScopedEnvExtension α β σ) (env : Lean.Environment) (b : β) :
                        Equations
                        Instances For
                          def Lean.ScopedEnvExtension.addScopedEntry {α : Type} {β : Type} {σ : Type} (ext : Lean.ScopedEnvExtension α β σ) (env : Lean.Environment) (namespaceName : Lake.Name) (b : β) :
                          Equations
                          Instances For
                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              def Lean.ScopedEnvExtension.addCore {α : Type} {β : Type} {σ : Type} (env : Lean.Environment) (ext : Lean.ScopedEnvExtension α β σ) (b : β) (kind : Lean.AttributeKind) (namespaceName : Lake.Name) :
                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                Equations
                                Instances For
                                  def Lean.ScopedEnvExtension.getState {σ : Type} {α : Type} {β : Type} [Inhabited σ] (ext : Lean.ScopedEnvExtension α β σ) (env : Lean.Environment) :
                                  σ
                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    def Lean.ScopedEnvExtension.activateScoped {α : Type} {β : Type} {σ : Type} (ext : Lean.ScopedEnvExtension α β σ) (env : Lean.Environment) (namespaceName : Lake.Name) :
                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      def Lean.ScopedEnvExtension.modifyState {α : Type} {β : Type} {σ : Type} (ext : Lean.ScopedEnvExtension α β σ) (env : Lean.Environment) (f : σσ) :
                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            def Lean.activateScoped {m : TypeType} [Monad m] [Lean.MonadEnv m] [MonadLiftT (ST IO.RealWorld) m] (namespaceName : Lake.Name) :
                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              @[reducible, inline]
                                              Equations
                                              Instances For
                                                • name : Lake.Name
                                                • addEntry : σασ
                                                • initial : σ
                                                • finalizeImport : σσ
                                                Instances For
                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For