Introduction

This is the rendered supplement to the paper "‘do’ Unchained: Embracing Local Imperativity in a Purely Functional Language". It contains

  • a reference implementation of the described do extensions using the Lean 4 macro system (lightly commented)
  • a Lean formalization of the translation function, the static and dynamic semantics, and the equivalence proof described in the paper in literate style

Basic do notation

open Lean

declare_syntax_cat stmt
syntax "do'" stmt : term

-- Prevent `if ...` from being parsed as a term
syntax (priority := low) term : stmt
syntax "let" ident "←" stmt:1 ";" stmt : stmt
macro "{" 
s: TSyntax `stmt
s
:stmt "}" : stmt => `($
s: TSyntax `stmt
s
) /- Remark: we annotate `macro`s and `macro_rules` with their corresponding translation/abbreviation id from the paper. -/ syntax "d!" stmt : term -- corresponds to `D(s)` macro_rules | `(do' $
s: TSyntax `stmt
s
) => `(d! $
s: TSyntax `stmt
s
) -- (1) -- helper function; see usage below def
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
(
s: TSyntax `stmt
s
:
TSyntax: SyntaxNodeKinds → Type
TSyntax
`stmt: Name
`stmt
) :
MacroM: Type → Type
MacroM
(
TSyntax: SyntaxNodeKinds → Type
TSyntax
`stmt: Name
`stmt
) := do let
s': Syntax
s'
expandMacros: Syntax → (optParam (SyntaxNodeKind → Bool) fun k => k != `Lean.Parser.Term.byTactic) → MacroM Syntax
expandMacros
s: TSyntax `stmt
s
if
s: TSyntax `stmt
s
==
s': Syntax
s'
then
Macro.throwUnsupported: {α : Type} → MacroM α
Macro.throwUnsupported
else -- There is no static guarantee that `expandMacros` stays in the `stmt` category, -- but it is true for all our macros return
TSyntax.mk: {ks : SyntaxNodeKinds} → Syntax → TSyntax ks
TSyntax.mk
s': Syntax
s'
macro_rules | `(d! $
e: TSyntax `term
e
:term) => `($
e: TSyntax `term
e
) -- (D1) | `(d! let $
x: TSyntax `ident
x
$
s: TSyntax `stmt
s
; $
s': TSyntax `stmt
s'
) => `((d! $
s: TSyntax `stmt
s
) >>= fun $
x: TSyntax `ident
x
=> (d! $
s': TSyntax `stmt
s'
)) -- (D2) | `(d! $
s: TSyntax `stmt
s
) => do -- fallback rule: try to expand abbreviation let
s': TSyntax `stmt
s'
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
s: TSyntax `stmt
s
`(d! $
s': TSyntax `stmt
s'
) macro "let"
x: TSyntax `ident
x
:ident ":="
e: TSyntax `term
e
:term ";"
s: TSyntax `stmt
s
:stmt : stmt => `(let $
x: TSyntax `ident
x
pure $
e: TSyntax `term
e
; $
s: TSyntax `stmt
s
) -- (A1) -- priority `0` prevents `;` from being used in trailing contexts without braces (see e.g. `:1` above) macro:0
s₁: TSyntax `stmt
s₁
:stmt ";"
s₂: TSyntax `stmt
s₂
:stmt : stmt => `(let x $
s₁: TSyntax `stmt
s₁
; $
s₂: TSyntax `stmt
s₂
) -- (A2)
import Do.Basic
import Do.LazyList

Local Mutation

open Lean

Disable the automatic monadic lifting feature described in the paper. We want to make it clear that we do not depend on it.

set_option autoLift false

syntax "let" "mut" ident ":=" term ";" stmt : stmt
syntax ident ":=" term : stmt
syntax "if" term "then" stmt "else" stmt:1 : stmt

declare_syntax_cat expander
-- generic syntax for traversal-like functions S_y/R/B/L
syntax "expand!" expander "in" stmt:1 : stmt
syntax "mut" ident : expander  -- corresponds to `S_y`

-- generic traversal rules
macro_rules
  | `(stmt| expand! $
exp: TSyntax `expander
exp
in let $
x: TSyntax `ident
x
$
s: TSyntax `stmt
s
; $
s': TSyntax `stmt
s'
) => `(stmt| let $
x: TSyntax `ident
x
expand! $
exp: TSyntax `expander
exp
in $
s: TSyntax `stmt
s
; expand! $
exp: TSyntax `expander
exp
in $
s': TSyntax `stmt
s'
) -- subsumes (R3, B4, L4) | `(stmt| expand! $
exp: TSyntax `expander
exp
in let mut $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; $
s': TSyntax `stmt
s'
) => `(stmt| let mut $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; expand! $
exp: TSyntax `expander
exp
in $
s': TSyntax `stmt
s'
) -- subsumes (R4, B5, L5) | `(stmt| expand! $_ in $
x: TSyntax `ident
x
:ident := $
e: TSyntax `term
e
) => `(stmt| $
x: TSyntax `ident
x
:ident := $
e: TSyntax `term
e
) -- subsumes (R5, B6, L6) | `(stmt| expand! $
exp: TSyntax `expander
exp
in if $
e: TSyntax `term
e
then $
s₁: TSyntax `stmt
s₁
else $
s₂: TSyntax `stmt
s₂
) => `(stmt| if $
e: TSyntax `term
e
then expand! $
exp: TSyntax `expander
exp
in $
s₁: TSyntax `stmt
s₁
else expand! $
exp: TSyntax `expander
exp
in $
s₂: TSyntax `stmt
s₂
) -- subsumes (S6, R6, B7, L7) | `(stmt| expand! $
exp: TSyntax `expander
exp
in $
s: TSyntax `stmt
s
) => do let
s': TSyntax `stmt
s'
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
s: TSyntax `stmt
s
`(stmt| expand! $
exp: TSyntax `expander
exp
in $
s': TSyntax `stmt
s'
) macro_rules | `(d! let mut $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; $
s: TSyntax `stmt
s
) => `(let $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; StateT.run' (d! expand! mut $
x: TSyntax `ident
x
in $
s: TSyntax `stmt
s
) $
x: TSyntax `ident
x
) -- (D3) | `(d! $
x: TSyntax `ident
x
:ident := $_:term) => -- `s!"..."` is an interpolated string. For more information, see https://leanprover.github.io/lean4/doc/stringinterp.html.
throw: {ε : Type} → {m : Type → Type} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
<|
Macro.Exception.error: Syntax → String → Macro.Exception
Macro.Exception.error
x: TSyntax `ident
x
s!"variable '{
x: TSyntax `ident
x
.
getId: Ident → Name
getId
}' is not reassignable in this scope" | `(d! if $
e: TSyntax `term
e
then $
s₁: TSyntax `stmt
s₁
else $
s₂: TSyntax `stmt
s₂
) => `(if $
e: TSyntax `term
e
then d! $
s₁: TSyntax `stmt
s₁
else d! $
s₂: TSyntax `stmt
s₂
) -- (D4) macro_rules | `(stmt| expand! mut $_ in $
e: TSyntax `term
e
:term) => `(stmt| StateT.lift $
e: TSyntax `term
e
) -- (S1) | `(stmt| expand! mut $
y: TSyntax `ident
y
in let $
x: TSyntax `ident
x
$
s: TSyntax `stmt
s
; $
s': TSyntax `stmt
s'
) => -- (S2) if
x: TSyntax `ident
x
==
y: TSyntax `ident
y
then
throw: {ε : Type} → {m : Type → Type} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
<|
Macro.Exception.error: Syntax → String → Macro.Exception
Macro.Exception.error
x: TSyntax `ident
x
s!"cannot shadow 'mut' variable '{
x: TSyntax `ident
x
.
getId: Ident → Name
getId
}'" else `(stmt| let $
x: TSyntax `ident
x
expand! mut $
y: TSyntax `ident
y
in $
s: TSyntax `stmt
s
; let $
y: TSyntax `ident
y
get; expand! mut $
y: TSyntax `ident
y
in $
s': TSyntax `stmt
s'
) | `(stmt| expand! mut $
y: TSyntax `ident
y
in let mut $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; $
s': TSyntax `stmt
s'
) => -- (S3) if
x: TSyntax `ident
x
==
y: TSyntax `ident
y
then
throw: {ε : Type} → {m : Type → Type} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
<|
Macro.Exception.error: Syntax → String → Macro.Exception
Macro.Exception.error
x: TSyntax `ident
x
s!"cannot shadow 'mut' variable '{
x: TSyntax `ident
x
.
getId: Ident → Name
getId
}'" else `(stmt| let mut $
x: TSyntax `ident
x
:= $
e: TSyntax `term
e
; expand! mut $
y: TSyntax `ident
y
in $
s': TSyntax `stmt
s'
) | `(stmt| expand! mut $
y: TSyntax `ident
y
in $
x: TSyntax `ident
x
:ident := $
e: TSyntax `term
e
) => if
x: TSyntax `ident
x
==
y: TSyntax `ident
y
then `(stmt| set $
e: TSyntax `term
e
) -- (S5) else `(stmt| $
x: TSyntax `ident
x
:ident := $
e: TSyntax `term
e
) -- (S4) macro:0 "let" "mut"
x: TSyntax `ident
x
:ident "←"
s: TSyntax `stmt
s
:stmt:1 ";"
s': TSyntax `stmt
s'
:stmt : stmt => `(let y $
s: TSyntax `stmt
s
; let mut $
x: TSyntax `ident
x
:= y; $
s': TSyntax `stmt
s'
) -- (A3) macro:0
x: TSyntax `ident
x
:ident "←"
s: TSyntax `stmt
s
:stmt:1 : stmt => `(let y $
s: TSyntax `stmt
s
; $
x: TSyntax `ident
x
:ident := y) -- (A4) -- a variant of (A4) since we technically cannot make the above macro a `stmt` macro:0
x: TSyntax `ident
x
:ident "←"
s: TSyntax `stmt
s
:stmt:1 ";"
s': TSyntax `stmt
s'
:stmt : stmt => `(let y $
s: TSyntax `stmt
s
; $
x: TSyntax `ident
x
:ident := y; $
s': TSyntax `stmt
s'
) macro "if"
e: TSyntax `term
e
:term "then"
s₁: TSyntax `stmt
s₁
:stmt:1 : stmt => `(if $
e: TSyntax `term
e
then $
s₁: TSyntax `stmt
s₁
else pure ()) -- (A5) macro "unless"
e: TSyntax `term
e
:term "do'"
s₂: TSyntax `stmt
s₂
:stmt:1 : stmt => `(if $
e: TSyntax `term
e
then pure () else $
s₂: TSyntax `stmt
s₂
) -- (A6) /- The `variable` command instructs Lean to insert the declared variables as bound variables in definitions that refer to them. -/ variable [
Monad: (Type → Type u_1) → Type (max 1 u_1)
Monad
m: Type u_1 → Type u_2
m
] variable (
ma: m α
ma
ma': m α
ma'
:
m: Type u_1 → Type u_2
m
α: Type ?u.10009
α
) /- Mark `map_eq_pure_bind` as a simplification lemma. It is a theorem for `f <$> x = x >>= pure (f a)` -/ attribute [local simp]
map_eq_pure_bind: ∀ {m : Type u_1 → Type u_2} {α β : Type u_1} [inst : Monad m] [inst_1 : LawfulMonad m] (f : α → β) (x : m α), f <$> x = do let a ← x pure (f a)
map_eq_pure_bind
/- Remark: an `example` in Lean is like a "nameless" definition, and it does not update the environment. It is useful for writing tests. -/ /- The instance `[LawfulMonad m]` contains the monadic laws. For more information, see https://github.com/leanprover/lean4/blob/v4.0.0-m4/src/Init/Control/Lawful.lean -/
example: ∀ {m : Type u_1 → Type u_2} {α : Type u_1} [inst : Monad m] (ma : m α) [inst_1 : LawfulMonad m], (do let y ← ma let x : α := y StateT.run' (StateT.lift (pure x)) x) = ma
example
[
LawfulMonad: (m : Type u_1 → Type u_2) → [inst : Monad m] → Prop
LawfulMonad
m: Type u_1 → Type u_2
m
] : (do' let mut
x: α
x
ma: m α
ma
;
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
x: α
x
:
m: Type u_1 → Type u_2
m
α: Type u_1
α
) =
ma: m α
ma
:=
m: Type ?u.10102 Type ?u.10058
α: Type ?u.10102
inst✝¹: Monad m
ma, ma': m α

(do let y ma let x : α := y StateT.run' (StateT.lift (pure x)) x) = ma

Goals accomplished! 🐙
example: ∀ {m : Type u_1 → Type u_2} {α : Type u_1} [inst : Monad m] (ma ma' : m α) [inst_1 : LawfulMonad m], (do let y ← ma let x : α := y StateT.run' (do let y ← StateT.lift ma' let _ ← get set y let x ← get StateT.lift (pure x)) x) = do let _ ← ma ma'
example
[
LawfulMonad: (m : Type u_1 → Type u_2) → [inst : Monad m] → Prop
LawfulMonad
m: Type u_1 → Type u_2
m
] : (do' let mut
x: α
x
ma: m α
ma
; x
ma': m α
ma'
;
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
x: α
x
) = (
ma: m α
ma
>>= fun _ =>
ma': m α
ma'
) :=
m: Type ?u.10970 Type ?u.10750
α: Type ?u.10970
inst✝¹: Monad m
ma, ma': m α

(do let y ma let x : α := y StateT.run' (do let y StateT.lift ma' let _ get set y let x get StateT.lift (pure x)) x) = do let _ ma ma'

Goals accomplished! 🐙
/- The command `#check_failure <term>` succeeds only if `<term>` fails to be elaborated. -/
do let y ma let x : α := y sorryAx (m ?m.13198) true : m ?m.13198
do let y ma let x : α := y sorryAx (m ?m.13198) true : m ?m.13198
Warning: cannot shadow 'mut' variable 'x'
do let y ma let x : α := y sorryAx (m ?m.13198) true : m ?m.13198
do let _ ma sorryAx (m ?m.13260) true : m ?m.13260
do let _ ma sorryAx (m ?m.13260) true : m ?m.13260
Warning: variable 'x' is not reassignable in this scope
do let _ ma sorryAx (m ?m.13260) true : m ?m.13260
variable (
b: Bool
b
:
Bool: Type
Bool
) -- The following equivalence is true even if `m` does not satisfy the monadic laws.
example: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (ma : m α) (b : Bool), (if b = true then discard ma else pure ()) = if b = true then discard ma else pure ()
example
: (do' if
b: Bool
b
then {
discard: {f : Type → Type u_1} → {α : Type} → [inst : Functor f] → f α → f PUnit
discard
ma: m α
ma
}) = (if
b: Bool
b
then
discard: {f : Type → Type u_1} → {α : Type} → [inst : Functor f] → f α → f PUnit
discard
ma: m α
ma
else
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
()) :=
rfl: ∀ {α : Type u_1} {a : α}, a = a
rfl
theorem
simple: ∀ [inst : LawfulMonad m], (do let y ← ma let x : α := y StateT.run' (do if b = true then do let y ← StateT.lift ma' let _ ← get set y else StateT.lift (pure ()) let x ← get StateT.lift (pure x)) x) = do let x ← ma if b = true then ma' else pure x
simple
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] : (do' let mut
x: α
x
ma: m α
ma
; if
b: Bool
b
then { x
ma': m α
ma'
};
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
x: α
x
) = (
ma: m α
ma
>>= fun
x: α
x
=> if
b: Bool
b
then
ma': m α
ma'
else
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
x: α
x
) :=
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool

(do let y ma let x : α := y StateT.run' (do if b = true then do let y StateT.lift ma' let _ get set y else StateT.lift (pure ()) let x get StateT.lift (pure x)) x) = do let x ma if b = true then ma' else pure x
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α

false
(do let y ma let x : α := y StateT.run' (do if false = true then do let y StateT.lift ma' let _ get set y else StateT.lift (pure ()) let x get StateT.lift (pure x)) x) = do let x ma if false = true then ma' else pure x
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
(do let y ma let x : α := y StateT.run' (do if true = true then do let y StateT.lift ma' let _ get set y else StateT.lift (pure ()) let x get StateT.lift (pure x)) x) = do let x ma if true = true then ma' else pure x
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α

false
(do let y ma let x : α := y StateT.run' (do if false = true then do let y StateT.lift ma' let _ get set y else StateT.lift (pure ()) let x get StateT.lift (pure x)) x) = do let x ma if false = true then ma' else pure x
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
(do let y ma let x : α := y StateT.run' (do if true = true then do let y StateT.lift ma' let _ get set y else StateT.lift (pure ()) let x get StateT.lift (pure x)) x) = do let x ma if true = true then ma' else pure x

Goals accomplished! 🐙
example: ∀ {m : Type u_1 → Type u_2} {α : Type u_1} [inst : Monad m] (ma ma' : m α) (b : Bool) [inst_1 : LawfulMonad m] (f : α → α → α), (do let y ← ma let x : α := y StateT.run' (do let y ← if b = true then do let y ← StateT.lift ma let _ ← get set y let _ ← get StateT.lift ma' else StateT.lift ma' let x ← get StateT.lift (pure (f x y))) x) = do let x ← ma if b = true then do let x ← ma let y ← ma' pure (f x y) else do let y ← ma' pure (f x y)
example
[
LawfulMonad: (m : Type u_1 → Type u_2) → [inst : Monad m] → Prop
LawfulMonad
m: Type u_1 → Type u_2
m
] (
f: α → α → α
f
:
α: Type u_1
α
α: Type u_1
α
α: Type u_1
α
) : (do' let mut
x: α
x
ma: m α
ma
; let
y: α
y
if
b: Bool
b
then { x
ma: m α
ma
;
ma': m α
ma'
} else {
ma': m α
ma'
};
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
(
f: α → α → α
f
x: α
x
y: α
y
)) = (
ma: m α
ma
>>= fun
x: α
x
=> if
b: Bool
b
then
ma: m α
ma
>>= fun
x: α
x
=>
ma': m α
ma'
>>= fun
y: α
y
=>
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
(
f: α → α → α
f
x: α
x
y: α
y
) else
ma': m α
ma'
>>= fun
y: α
y
=>
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
(
f: α → α → α
f
x: α
x
y: α
y
)) :=
m: Type ?u.17614 Type ?u.17083
α: Type ?u.17614
inst✝¹: Monad m
ma, ma': m α
b: Bool
f: α α α

(do let y ma let x : α := y StateT.run' (do let y if b = true then do let y StateT.lift ma let _ get set y let _ get StateT.lift ma' else StateT.lift ma' let x get StateT.lift (pure (f x y))) x) = do let x ma if b = true then do let x ma let y ma' pure (f x y) else do let y ma' pure (f x y)
m: Type ?u.17614 Type ?u.17083
α: Type ?u.17614
inst✝¹: Monad m
ma, ma': m α
f: α α α

false
(do let y ma let x : α := y StateT.run' (do let y if false = true then do let y StateT.lift ma let _ get set y let _ get StateT.lift ma' else StateT.lift ma' let x get StateT.lift (pure (f x y))) x) = do let x ma if false = true then do let x ma let y ma' pure (f x y) else do let y ma' pure (f x y)
m: Type ?u.17614 Type ?u.17083
α: Type ?u.17614
inst✝¹: Monad m
ma, ma': m α
f: α α α
(do let y ma let x : α := y StateT.run' (do let y if true = true then do let y StateT.lift ma let _ get set y let _ get StateT.lift ma' else StateT.lift ma' let x get StateT.lift (pure (f x y))) x) = do let x ma if true = true then do let x ma let y ma' pure (f x y) else do let y ma' pure (f x y)
m: Type ?u.17614 Type ?u.17083
α: Type ?u.17614
inst✝¹: Monad m
ma, ma': m α
f: α α α

false
(do let y ma let x : α := y StateT.run' (do let y if false = true then do let y StateT.lift ma let _ get set y let _ get StateT.lift ma' else StateT.lift ma' let x get StateT.lift (pure (f x y))) x) = do let x ma if false = true then do let x ma let y ma' pure (f x y) else do let y ma' pure (f x y)
m: Type ?u.17614 Type ?u.17083
α: Type ?u.17614
inst✝¹: Monad m
ma, ma': m α
f: α α α
(do let y ma let x : α := y StateT.run' (do let y if true = true then do let y StateT.lift ma let _ get set y let _ get StateT.lift ma' else StateT.lift ma' let x get StateT.lift (pure (f x y))) x) = do let x ma if true = true then do let x ma let y ma' pure (f x y) else do let y ma' pure (f x y)

Goals accomplished! 🐙
/- Nondeterminism example from Section 2. -/ def
choose: {α : Type u_1} → List α → LazyList α
choose
:= @
List.toLazy: {α : Type u_1} → List α → LazyList α
List.toLazy
def
ex: LazyList Nat
ex
:
LazyList: Type → Type
LazyList
Nat: Type
Nat
:= do' let mut
x: Nat
x
:=
0: Nat
0
; let
y: Nat
y
choose: {α : Type} → List α → LazyList α
choose
[
0: Nat
0
,
1: Nat
1
,
2: Nat
2
,
3: Nat
3
]; x :=
x: Nat
x
+
1: Nat
1
;
guard: {f : Type → Type} → [inst : Alternative f] → (p : Prop) → [inst : Decidable p] → f Unit
guard
(
x: Nat
x
<
3: Nat
3
);
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
(
x: Nat
x
+
y: Nat
y
) -- Generate all solutions
[1, 2, 3, 4]
ex: LazyList Nat
ex
.
toList: {α : Type} → LazyList α → List α
toList

LazyList Nondeterminism Monad

inductive 
LazyList: Type u → Type u
LazyList
(
α: Type u
α
:
Type u: Type (u + 1)
Type u
) where |
nil: {α : Type u} → LazyList α
nil
:
LazyList: Type u → Type u
LazyList
α: Type u
α
|
cons: {α : Type u} → α → LazyList α → LazyList α
cons
(
hd: α
hd
:
α: Type u
α
) (
tl: LazyList α
tl
:
LazyList: Type u → Type u
LazyList
α: Type u
α
) :
LazyList: Type u → Type u
LazyList
α: Type u
α
|
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
t: Thunk (LazyList α)
t
:
Thunk: Type u → Type u
Thunk
$
LazyList: Type u → Type u
LazyList
α: Type u
α
) :
LazyList: Type u → Type u
LazyList
α: Type u
α
def
List.toLazy: {α : Type u} → List α → LazyList α
List.toLazy
{
α: Type u
α
:
Type u: Type (u + 1)
Type u
} :
List: Type u → Type u
List
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
| [] =>
LazyList.nil: {α : Type u} → LazyList α
LazyList.nil
|
h: α
h
::
t: List α
t
=>
LazyList.cons: {α : Type u} → α → LazyList α → LazyList α
LazyList.cons
h: α
h
(
toLazy: {α : Type u} → List α → LazyList α
toLazy
t: List α
t
) namespace LazyList variable {
α: Type u
α
:
Type u: Type (u + 1)
Type u
} {
β: Type v
β
:
Type v: Type (v + 1)
Type v
} {
δ: Type w
δ
:
Type w: Type (w + 1)
Type w
}
instance: {α : Type u} → Inhabited (LazyList α)
instance
:
Inhabited: Type u → Type u
Inhabited
(
LazyList: Type u → Type u
LazyList
α: Type u
α
) := ⟨
nil: {α : Type u} → LazyList α
nil
@[inline] def
pure: α → LazyList α
pure
:
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
a: α
a
=>
cons: {α : Type u} → α → LazyList α → LazyList α
cons
a: α
a
nil: {α : Type u} → LazyList α
nil
partial def
isEmpty: LazyList α → Bool
isEmpty
:
LazyList: Type u → Type u
LazyList
α: Type u
α
Bool: Type
Bool
|
nil: {α : Type ?u.1027} → LazyList α
nil
=>
true: Bool
true
|
cons: {α : Type ?u.1036} → α → LazyList α → LazyList α
cons
_ _ =>
false: Bool
false
|
delayed: {α : Type ?u.1058} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
isEmpty: LazyList α → Bool
isEmpty
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
partial def
toList: LazyList α → List α
toList
:
LazyList: Type u → Type u
LazyList
α: Type u
α
List: Type u → Type u
List
α: Type u
α
|
nil: {α : Type ?u.1187} → LazyList α
nil
=>
[]: List α
[]
|
cons: {α : Type ?u.1199} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
=>
a: α
a
::
toList: LazyList α → List α
toList
as: LazyList α
as
|
delayed: {α : Type ?u.1221} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
toList: LazyList α → List α
toList
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
partial def
head: {α : Type u} → [inst : Inhabited α] → LazyList α → α
head
[
Inhabited: Type u → Type u
Inhabited
α: Type u
α
] :
LazyList: Type u → Type u
LazyList
α: Type u
α
α: Type u
α
|
nil: {α : Type ?u.1357} → LazyList α
nil
=>
default: {α : Type u} → [self : Inhabited α] → α
default
|
cons: {α : Type ?u.1373} → α → LazyList α → LazyList α
cons
a: α
a
Warning: unused variable `as`
=>
a: α
a
|
delayed: {α : Type ?u.1393} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
head: [inst : Inhabited α] → LazyList α → α
head
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
partial def
tail: {α : Type u} → LazyList α → LazyList α
tail
:
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
nil: {α : Type ?u.1530} → LazyList α
nil
=>
nil: {α : Type u} → LazyList α
nil
|
cons: {α : Type ?u.1541} → α → LazyList α → LazyList α
cons
Warning: unused variable `a`
as: LazyList α
as
=>
as: LazyList α
as
|
delayed: {α : Type ?u.1561} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
tail: LazyList α → LazyList α
tail
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
partial def
append: {α : Type u} → LazyList α → (Unit → LazyList α) → LazyList α
append
:
LazyList: Type u → Type u
LazyList
α: Type u
α
(
Unit: Type
Unit
LazyList: Type u → Type u
LazyList
α: Type u
α
)
LazyList: Type u → Type u
LazyList
α: Type u
α
|
nil: {α : Type ?u.1704} → LazyList α
nil
,
bs: Unit → LazyList α
bs
=>
bs: Unit → LazyList α
bs
(): Unit
()
|
cons: {α : Type ?u.1723} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
,
bs: Unit → LazyList α
bs
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
cons: {α : Type u} → α → LazyList α → LazyList α
cons
a: α
a
(
append: LazyList α → (Unit → LazyList α) → LazyList α
append
as: LazyList α
as
bs: Unit → LazyList α
bs
)) |
delayed: {α : Type ?u.1783} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
,
bs: Unit → LazyList α
bs
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
append: LazyList α → (Unit → LazyList α) → LazyList α
append
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
bs: Unit → LazyList α
bs
)
instance: {α : Type u} → Append (LazyList α)
instance
:
Append: Type u → Type u
Append
(
LazyList: Type u → Type u
LazyList
α: Type u
α
) where append
a: LazyList α
a
b: LazyList α
b
:=
LazyList.append: {α : Type u} → LazyList α → (Unit → LazyList α) → LazyList α
LazyList.append
a: LazyList α
a
(fun _ =>
b: LazyList α
b
) partial def
interleave: LazyList α → LazyList α → LazyList α
interleave
:
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
nil: {α : Type ?u.2013} → LazyList α
nil
,
bs: LazyList α
bs
=>
bs: LazyList α
bs
|
cons: {α : Type ?u.2031} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
,
bs: LazyList α
bs
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
cons: {α : Type u} → α → LazyList α → LazyList α
cons
a: α
a
(
interleave: LazyList α → LazyList α → LazyList α
interleave
bs: LazyList α
bs
as: LazyList α
as
)) |
delayed: {α : Type ?u.2089} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
,
bs: LazyList α
bs
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
interleave: LazyList α → LazyList α → LazyList α
interleave
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
bs: LazyList α
bs
) partial def
map: (α → β) → LazyList α → LazyList β
map
(
f: α → β
f
:
α: Type u
α
β: Type v
β
) :
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type v → Type v
LazyList
β: Type v
β
|
nil: {α : Type ?u.2270} → LazyList α
nil
=>
nil: {α : Type v} → LazyList α
nil
|
cons: {α : Type ?u.2282} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
=>
delayed: {α : Type v} → Thunk (LazyList α) → LazyList α
delayed
(
cons: {α : Type v} → α → LazyList α → LazyList α
cons
(
f: α → β
f
a: α
a
) (
map: (α → β) → LazyList α → LazyList β
map
f: α → β
f
as: LazyList α
as
)) |
delayed: {α : Type ?u.2332} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
delayed: {α : Type v} → Thunk (LazyList α) → LazyList α
delayed
(
map: (α → β) → LazyList α → LazyList β
map
f: α → β
f
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
) partial def
map₂: (α → β → δ) → LazyList α → LazyList β → LazyList δ
map₂
(
f: α → β → δ
f
:
α: Type u
α
β: Type v
β
δ: Type w
δ
) :
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type v → Type v
LazyList
β: Type v
β
LazyList: Type w → Type w
LazyList
δ: Type w
δ
|
nil: {α : Type ?u.2498} → LazyList α
nil
, _ =>
nil: {α : Type w} → LazyList α
nil
| _,
nil: {α : Type ?u.2518} → LazyList α
nil
=>
nil: {α : Type w} → LazyList α
nil
|
cons: {α : Type ?u.2536} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
,
cons: {α : Type ?u.2540} → α → LazyList α → LazyList α
cons
b: β
b
bs: LazyList β
bs
=>
delayed: {α : Type w} → Thunk (LazyList α) → LazyList α
delayed
(
cons: {α : Type w} → α → LazyList α → LazyList α
cons
(
f: α → β → δ
f
a: α
a
b: β
b
) (
map₂: (α → β → δ) → LazyList α → LazyList β → LazyList δ
map₂
f: α → β → δ
f
as: LazyList α
as
bs: LazyList β
bs
)) |
delayed: {α : Type ?u.2607} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
,
bs: LazyList β
bs
=>
delayed: {α : Type w} → Thunk (LazyList α) → LazyList α
delayed
(
map₂: (α → β → δ) → LazyList α → LazyList β → LazyList δ
map₂
f: α → β → δ
f
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
bs: LazyList β
bs
) |
as: LazyList α
as
,
delayed: {α : Type ?u.2644} → Thunk (LazyList α) → LazyList α
delayed
bs: Thunk (LazyList β)
bs
=>
delayed: {α : Type w} → Thunk (LazyList α) → LazyList α
delayed
(
map₂: (α → β → δ) → LazyList α → LazyList β → LazyList δ
map₂
f: α → β → δ
f
as: LazyList α
as
bs: Thunk (LazyList β)
bs
.
get: {α : Type v} → Thunk α → α
get
) @[inline] def
zip: LazyList α → LazyList β → LazyList (α × β)
zip
:
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type v → Type v
LazyList
β: Type v
β
LazyList: Type (max v u) → Type (max v u)
LazyList
(
α: Type u
α
×
β: Type v
β
) :=
map₂: {α : Type u} → {β : Type v} → {δ : Type (max u v)} → (α → β → δ) → LazyList α → LazyList β → LazyList δ
map₂
Prod.mk: {α : Type u} → {β : Type v} → α → β → α × β
Prod.mk
partial def
join: LazyList (LazyList α) → LazyList α
join
:
LazyList: Type u → Type u
LazyList
(
LazyList: Type u → Type u
LazyList
α: Type u
α
)
LazyList: Type u → Type u
LazyList
α: Type u
α
|
nil: {α : Type ?u.3036} → LazyList α
nil
=>
nil: {α : Type u} → LazyList α
nil
|
cons: {α : Type ?u.3047} → α → LazyList α → LazyList α
cons
a: LazyList α
a
as: LazyList (LazyList α)
as
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
append: {α : Type u} → LazyList α → (Unit → LazyList α) → LazyList α
append
a: LazyList α
a
fun _ =>
join: LazyList (LazyList α) → LazyList α
join
as: LazyList (LazyList α)
as
) |
delayed: {α : Type ?u.3102} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList (LazyList α))
as
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
join: LazyList (LazyList α) → LazyList α
join
as: Thunk (LazyList (LazyList α))
as
.
get: {α : Type u} → Thunk α → α
get
) @[inline] partial def
bind: {α : Type u} → {β : Type v} → LazyList α → (α → LazyList β) → LazyList β
bind
(
x: LazyList α
x
:
LazyList: Type u → Type u
LazyList
α: Type u
α
) (
f: α → LazyList β
f
:
α: Type u
α
LazyList: Type v → Type v
LazyList
β: Type v
β
) :
LazyList: Type v → Type v
LazyList
β: Type v
β
:=
join: {α : Type v} → LazyList (LazyList α) → LazyList α
join
(
x: LazyList α
x
.
map: {α : Type u} → {β : Type v} → (α → β) → LazyList α → LazyList β
map
f: α → LazyList β
f
)
instance: Monad LazyList
instance
:
Monad: (Type u_1 → Type u_1) → Type (u_1 + 1)
Monad
LazyList: Type u_1 → Type u_1
LazyList
where pure :=
LazyList.pure: {α : Type u_1} → α → LazyList α
LazyList.pure
bind :=
LazyList.bind: {α β : Type u_1} → LazyList α → (α → LazyList β) → LazyList β
LazyList.bind
map :=
LazyList.map: {α β : Type u_1} → (α → β) → LazyList α → LazyList β
LazyList.map
instance: Alternative LazyList
instance
:
Alternative: (Type u_1 → Type u_1) → Type (u_1 + 1)
Alternative
LazyList: Type u_1 → Type u_1
LazyList
where failure :=
nil: {α : Type u_1} → LazyList α
nil
orElse :=
LazyList.append: {α : Type u_1} → LazyList α → (Unit → LazyList α) → LazyList α
LazyList.append
partial def
approx: Nat → LazyList α → List α
approx
:
Nat: Type
Nat
LazyList: Type u → Type u
LazyList
α: Type u
α
List: Type u → Type u
List
α: Type u
α
|
0: Nat
0
,
Warning: unused variable `as`
=>
[]: List α
[]
| _,
nil: {α : Type ?u.3731} → LazyList α
nil
=>
[]: List α
[]
|
i: Nat
i
+1,
cons: {α : Type ?u.3749} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
=>
a: α
a
::
approx: Nat → LazyList α → List α
approx
i: Nat
i
as: LazyList α
as
|
i: Nat
i
+1,
delayed: {α : Type ?u.3839} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
approx: Nat → LazyList α → List α
approx
(
i: Nat
i
+
1: Nat
1
)
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
partial def
iterate: (α → α) → α → LazyList α
iterate
(
f: α → α
f
:
α: Type u
α
α: Type u
α
) :
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
x: α
x
=>
cons: {α : Type u} → α → LazyList α → LazyList α
cons
x: α
x
(
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
iterate: (α → α) → α → LazyList α
iterate
f: α → α
f
(
f: α → α
f
x: α
x
))) partial def
iterate₂: (α → α → α) → α → α → LazyList α
iterate₂
(
f: α → α → α
f
:
α: Type u
α
α: Type u
α
α: Type u
α
) :
α: Type u
α
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
x: α
x
,
y: α
y
=>
cons: {α : Type u} → α → LazyList α → LazyList α
cons
x: α
x
(
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
iterate₂: (α → α → α) → α → α → LazyList α
iterate₂
f: α → α → α
f
y: α
y
(
f: α → α → α
f
x: α
x
y: α
y
))) partial def
filter: (α → Bool) → LazyList α → LazyList α
filter
(
p: α → Bool
p
:
α: Type u
α
Bool: Type
Bool
) :
LazyList: Type u → Type u
LazyList
α: Type u
α
LazyList: Type u → Type u
LazyList
α: Type u
α
|
nil: {α : Type ?u.4392} → LazyList α
nil
=>
nil: {α : Type u} → LazyList α
nil
|
cons: {α : Type ?u.4403} → α → LazyList α → LazyList α
cons
a: α
a
as: LazyList α
as
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(if
p: α → Bool
p
a: α
a
then
cons: {α : Type u} → α → LazyList α → LazyList α
cons
a: α
a
(
filter: (α → Bool) → LazyList α → LazyList α
filter
p: α → Bool
p
as: LazyList α
as
) else
filter: (α → Bool) → LazyList α → LazyList α
filter
p: α → Bool
p
as: LazyList α
as
) |
delayed: {α : Type ?u.4545} → Thunk (LazyList α) → LazyList α
delayed
as: Thunk (LazyList α)
as
=>
delayed: {α : Type u} → Thunk (LazyList α) → LazyList α
delayed
(
filter: (α → Bool) → LazyList α → LazyList α
filter
p: α → Bool
p
as: Thunk (LazyList α)
as
.
get: {α : Type u} → Thunk α → α
get
) end LazyList
import Do.Mut

Early Return

open Lean

/- Disable the automatic monadic lifting feature described in the paper.
   We want to make it clear that we do not depend on it. -/
set_option autoLift false

def 
runCatch: {m : Type u_1 → Type u_2} → {α : Type u_1} → [inst : Monad m] → ExceptT α m α → m α
runCatch
[
Monad: (Type ?u.9 → Type ?u.8) → Type (max (?u.9 + 1) ?u.8)
Monad
m: Type u_1 → Type u_2
m
] (
x: ExceptT α m α
x
:
ExceptT: Type u_1 → (Type u_1 → Type u_2) → Type u_1 → Type u_2
ExceptT
α: Type u_1
α
m: Type u_1 → Type u_2
m
α: Type u_1
α
) :
m: Type u_1 → Type u_2
m
α: Type u_1
α
:=
ExceptT.run: {ε : Type u_1} → {m : Type u_1 → Type u_2} → {α : Type u_1} → ExceptT ε m α → m (Except ε α)
ExceptT.run
x: ExceptT α m α
x
>>= fun |
Except.ok: {ε : Type ?u.57} → {α : Type ?u.56} → α → Except ε α
Except.ok
x: α
x
=>
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
x: α
x
|
Except.error: {ε : Type ?u.115} → {α : Type ?u.114} → ε → Except ε α
Except.error
e: α
e
=>
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
e: α
e
/-- Count syntax nodes satisfying `p`. -/ partial def
Lean.Syntax.count: Syntax → (Syntax → Bool) → Nat
Lean.Syntax.count
(
stx: Syntax
stx
:
Syntax: Type
Syntax
) (
p: Syntax → Bool
p
:
Syntax: Type
Syntax
Bool: Type
Bool
) :
Nat: Type
Nat
:=
stx: Syntax
stx
.
getArgs: Syntax → Array Syntax
getArgs
.
foldl: {α β : Type} → (β → α → β) → β → (as : Array α) → optParam Nat 0 → optParam Nat (Array.size as) → β
foldl
(fun
n: Nat
n
arg: Syntax
arg
=>
n: Nat
n
+
arg: Syntax
arg
.
count: Syntax → (Syntax → Bool) → Nat
count
p: Syntax → Bool
p
) (if
p: Syntax → Bool
p
stx: Syntax
stx
then
1: Nat
1
else
0: Nat
0
) syntax "return" term : stmt syntax "return" : expander macro_rules | `(do' $
s: TSyntax `stmt
s
) => do -- (1') -- optimization: fall back to original rule (1) if now `return` statement was expanded let
s': TSyntax `stmt
s'
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
(← `(stmt| expand! return in $s)): TSyntax `stmt
( `(stmt|
(← `(stmt| expand! return in $s)): TSyntax `stmt
expand!
(← `(stmt| expand! return in $s)): TSyntax `stmt
(← `(stmt| expand! return in $s)): TSyntax `stmt
return
(← `(stmt| expand! return in $s)): TSyntax `stmt
(← `(stmt| expand! return in $s)): TSyntax `stmt
in
(← `(stmt| expand! return in $s)): TSyntax `stmt
$
s: TSyntax `stmt
s
(← `(stmt| expand! return in $s)): TSyntax `stmt
))
if
s': TSyntax `stmt
s'
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| return $_)) ==
s: TSyntax `stmt
s
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| return $_)) then `(d! $
s: TSyntax `stmt
s
) else `(ExceptCpsT.runCatch (d! $
s': TSyntax `stmt
s'
)) macro_rules | `(stmt| expand! return in return $
e: TSyntax `term
e
) => `(stmt| throw $
e: TSyntax `term
e
) -- (R1) | `(stmt| expand! return in $
e: TSyntax `term
e
:term) => `(stmt| ExceptCpsT.lift $
e: TSyntax `term
e
) -- (R2) variable [
Monad: (Type ?u.3135 → Type ?u.3134) → Type (max (?u.3135 + 1) ?u.3134)
Monad
m: Type ?u.2574 → Type ?u.2573
m
] variable (
ma: m α
ma
ma': m α
ma'
:
m: Type ?u.2585 → Type ?u.2584
m
α: Type ?u.2585
α
) variable (
b: Bool
b
:
Bool: Type
Bool
)
example: ∀ {m : Type u_1 → Type u_2} {α : Type u_1} [inst : Monad m] (ma : m α) [inst_1 : LawfulMonad m], (ExceptCpsT.runCatch do let x ← ExceptCpsT.lift ma throw x) = ma
example
[
LawfulMonad: (m : Type u_1 → Type u_2) → [inst : Monad m] → Prop
LawfulMonad
m: Type u_1 → Type u_2
m
] : (do' let
x: α
x
ma: m α
ma
; return
x: α
x
) =
ma: m α
ma
:=
m: Type ?u.2639 Type ?u.2638
α: Type ?u.2639
inst✝¹: Monad m
ma, ma': m α
b: Bool

(ExceptCpsT.runCatch do let x ExceptCpsT.lift ma throw x) = ma

Goals accomplished! 🐙
example: Id.run (ExceptCpsT.runCatch do let x ← ExceptCpsT.lift (pure 1) throw x) = 1
example
:
Id.run: {α : Type} → Id α → α
Id.run
(do' let
x: Nat
x
:=
1: Nat
1
; return
x: Nat
x
) =
1: Nat
1
:=
rfl: ∀ {α : Type} {a : α}, a = a
rfl
example: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (ma ma' : m α) (b : Bool) [inst_1 : LawfulMonad m], (ExceptCpsT.runCatch do if b = true then do let x ← ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if b = true then ma else ma'
example
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] : (do' if
b: Bool
b
then { let
x: α
x
ma: m α
ma
; return
x: α
x
};
ma': m α
ma'
) = (if
b: Bool
b
then
ma: m α
ma
else
ma': m α
ma'
) :=
m: Type Type ?u.3364
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool

(ExceptCpsT.runCatch do if b = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if b = true then ma else ma'
m: Type Type ?u.3364
α: Type
inst✝¹: Monad m
ma, ma': m α

false
(ExceptCpsT.runCatch do if false = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if false = true then ma else ma'
m: Type Type ?u.3364
α: Type
inst✝¹: Monad m
ma, ma': m α
(ExceptCpsT.runCatch do if true = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if true = true then ma else ma'
m: Type Type ?u.3364
α: Type
inst✝¹: Monad m
ma, ma': m α

false
(ExceptCpsT.runCatch do if false = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if false = true then ma else ma'
m: Type Type ?u.3364
α: Type
inst✝¹: Monad m
ma, ma': m α
(ExceptCpsT.runCatch do if true = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift (pure ()) ExceptCpsT.lift ma') = if true = true then ma else ma'

Goals accomplished! 🐙
example: ∀ {m : Type u_1 → Type u_2} {α : Type u_1} [inst : Monad m] (ma ma' : m α) (b : Bool) [inst_1 : LawfulMonad m], (ExceptCpsT.runCatch do let y ← if b = true then do let x ← ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if b = true then ma else ma'
example
[
LawfulMonad: (m : Type u_1 → Type u_2) → [inst : Monad m] → Prop
LawfulMonad
m: Type u_1 → Type u_2
m
] : (do' let
y: α
y
if
b: Bool
b
then { let
x: α
x
ma: m α
ma
; return
x: α
x
} else {
ma': m α
ma'
};
pure: {f : Type u_1 → Type u_2} → [self : Pure f] → {α : Type u_1} → α → f α
pure
y: α
y
) = (if
b: Bool
b
then
ma: m α
ma
else
ma': m α
ma'
) :=
m: Type ?u.4867 Type ?u.4866
α: Type ?u.4867
inst✝¹: Monad m
ma, ma': m α
b: Bool

(ExceptCpsT.runCatch do let y if b = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if b = true then ma else ma'
m: Type ?u.4867 Type ?u.4866
α: Type ?u.4867
inst✝¹: Monad m
ma, ma': m α

false
(ExceptCpsT.runCatch do let y if false = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if false = true then ma else ma'
m: Type ?u.4867 Type ?u.4866
α: Type ?u.4867
inst✝¹: Monad m
ma, ma': m α
(ExceptCpsT.runCatch do let y if true = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if true = true then ma else ma'
m: Type ?u.4867 Type ?u.4866
α: Type ?u.4867
inst✝¹: Monad m
ma, ma': m α

false
(ExceptCpsT.runCatch do let y if false = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if false = true then ma else ma'
m: Type ?u.4867 Type ?u.4866
α: Type ?u.4867
inst✝¹: Monad m
ma, ma': m α
(ExceptCpsT.runCatch do let y if true = true then do let x ExceptCpsT.lift ma throw x else ExceptCpsT.lift ma' ExceptCpsT.lift (pure y)) = if true = true then ma else ma'

Goals accomplished! 🐙
import Do.Return

Iteration

open Lean

/- Disable the automatic monadic lifting feature described in the paper.
   We want to make it clear that we do not depend on it. -/
set_option autoLift false

syntax "for" ident "in" term "do'" stmt:1 : stmt
syntax "break " : stmt
syntax "continue " : stmt

syntax "break" : expander
syntax "continue" : expander
syntax "lift" : expander

macro_rules
  | `(stmt| expand! $_   in break) => 
`(stmt| break): MacroM Syntax
`(stmt|
`(stmt| break): MacroM Syntax
break
`(stmt| break): MacroM Syntax
)
-- subsumes (S7, R7, B2, L1) | `(stmt| expand! $_ in continue) =>
`(stmt| continue): MacroM Syntax
`(stmt|
`(stmt| continue): MacroM Syntax
continue
`(stmt| continue): MacroM Syntax
)
-- subsumes (S8, R8, L2) | `(stmt| expand! $
exp: TSyntax `expander
exp
in for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' $
s: TSyntax `stmt
s
) => `(stmt| for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' expand! $
exp: TSyntax `expander
exp
in $
s: TSyntax `stmt
s
) -- subsumes (L8, R9) macro_rules | `(d! for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' $
s: TSyntax `stmt
s
) => do -- (D5), optimized like (1') let mut
s: TSyntax `stmt
s
:=
s: TSyntax `stmt
s
let
sb: TSyntax `stmt
sb
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
(← `(stmt| expand! break in $s)): TSyntax `stmt
( `(stmt|
(← `(stmt| expand! break in $s)): TSyntax `stmt
expand!
(← `(stmt| expand! break in $s)): TSyntax `stmt
(← `(stmt| expand! break in $s)): TSyntax `stmt
break
(← `(stmt| expand! break in $s)): TSyntax `stmt
(← `(stmt| expand! break in $s)): TSyntax `stmt
in
(← `(stmt| expand! break in $s)): TSyntax `stmt
$
s: TSyntax `stmt
s
(← `(stmt| expand! break in $s)): TSyntax `stmt
))
let
hasBreak: Prop
hasBreak
:=
sb: TSyntax `stmt
sb
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| break)) <
s: TSyntax `stmt
s
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| break)) if
hasBreak: Prop
hasBreak
then
s: TSyntax `stmt
s
:=
sb: TSyntax `stmt
sb
let
sc: TSyntax `stmt
sc
expandStmt: TSyntax `stmt → MacroM (TSyntax `stmt)
expandStmt
(← `(stmt| expand! continue in $s)): TSyntax `stmt
( `(stmt|
(← `(stmt| expand! continue in $s)): TSyntax `stmt
expand!
(← `(stmt| expand! continue in $s)): TSyntax `stmt
(← `(stmt| expand! continue in $s)): TSyntax `stmt
continue
(← `(stmt| expand! continue in $s)): TSyntax `stmt
(← `(stmt| expand! continue in $s)): TSyntax `stmt
in
(← `(stmt| expand! continue in $s)): TSyntax `stmt
$
s: TSyntax `stmt
s
(← `(stmt| expand! continue in $s)): TSyntax `stmt
))
let
hasContinue: Prop
hasContinue
:=
sc: TSyntax `stmt
sc
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| continue)) <
s: TSyntax `stmt
s
.
raw: {ks : SyntaxNodeKinds} → TSyntax ks → Syntax
raw
.
count: Syntax → (Syntax → Bool) → Nat
count
(· matches `(stmt| continue)) if
hasContinue: Prop
hasContinue
then
s: TSyntax `stmt
s
:=
sc: TSyntax `stmt
sc
let mut
body: TSyntax `term
body
`(d! $
s: TSyntax `stmt
s
) if
hasContinue: Prop
hasContinue
then
body: TSyntax `term
body
`(ExceptCpsT.runCatch $
body: TSyntax `term
body
) let mut
loop: TSyntax `term
loop
`(forM $
e: TSyntax `term
e
(fun $
x: TSyntax `ident
x
=> $
body: TSyntax `term
body
)) if
hasBreak: Prop
hasBreak
then
loop: TSyntax `term
loop
`(ExceptCpsT.runCatch $
loop: TSyntax `term
loop
)
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
loop: TSyntax `term
loop
| `(d! break%$
b: Syntax
b
) =>
throw: {ε : Type} → {m : Type → Type} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
<|
Macro.Exception.error: Syntax → String → Macro.Exception
Macro.Exception.error
b: Syntax
b
"unexpected 'break' outside loop": String
"unexpected 'break' outside loop"
| `(d! continue%$
c: Syntax
c
) =>
throw: {ε : Type} → {m : Type → Type} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
<|
Macro.Exception.error: Syntax → String → Macro.Exception
Macro.Exception.error
c: Syntax
c
"unexpected 'continue' outside loop": String
"unexpected 'continue' outside loop"
macro_rules | `(stmt| expand! break in break) =>
`(stmt| throw ()): MacroM Syntax
`(stmt| throw ())
-- (B1) | `(stmt| expand! break in $
e: TSyntax `term
e
:term) => `(stmt| ExceptCpsT.lift $
e: TSyntax `term
e
) -- (B3) | `(stmt| expand! break in for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' $
s: TSyntax `stmt
s
) => `(stmt| for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' expand! lift in $
s: TSyntax `stmt
s
) -- (B8) | `(stmt| expand! continue in continue) =>
`(stmt| throw ()): MacroM Syntax
`(stmt| throw ())
| `(stmt| expand! continue in $
e: TSyntax `term
e
:term) => `(stmt| ExceptCpsT.lift $
e: TSyntax `term
e
) | `(stmt| expand! continue in for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' $
s: TSyntax `stmt
s
) => `(stmt| for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' expand! lift in $
s: TSyntax `stmt
s
) macro_rules | `(stmt| expand! lift in $
e: TSyntax `term
e
:term) => `(stmt| ExceptCpsT.lift $
e: TSyntax `term
e
) -- (L3) macro_rules | `(stmt| expand! mut $
y: TSyntax `ident
y
in for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' $
s: TSyntax `stmt
s
) => `(stmt| for $
x: TSyntax `ident
x
in $
e: TSyntax `term
e
do' { let $
y: TSyntax `ident
y
get; expand! mut $
y: TSyntax `ident
y
in $
s: TSyntax `stmt
s
}) -- (S9) variable [
Monad: (Type ?u.36846 → Type ?u.36845) → Type (max (?u.36846 + 1) ?u.36845)
Monad
m: Type ?u.7311 → Type ?u.7310
m
] variable (
ma: m α
ma
ma': m α
ma'
:
m: Type ?u.7322 → Type ?u.7321
m
α: Type
α
) variable (
b: Bool
b
:
Bool: Type
Bool
) variable (
xs: List α
xs
:
List: Type → Type
List
α: Type
α
) (
act: α → m Unit
act
:
α: Type
α
m: Type → Type u_1
m
Unit: Type
Unit
) attribute [local simp]
map_eq_pure_bind: ∀ {m : Type u_1 → Type u_2} {α β : Type u_1} [inst : Monad m] [inst_1 : LawfulMonad m] (f : α → β) (x : m α), f <$> x = do let a ← x pure (f a)
map_eq_pure_bind
example: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (xs : List α) (act : α → m Unit) [inst_1 : LawfulMonad m], (forM xs fun x => act x) = List.forM xs act
example
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] : (do' for
x: α
x
in
xs: List α
xs
do' {
act: α → m Unit
act
x: α
x
}) =
xs: List α
xs
.
forM: {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → List α → (α → m PUnit) → m PUnit
forM
act: α → m Unit
act
:=
m: Type Type ?u.7430
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit

(forM xs fun x => act x) = List.forM xs act
m: Type Type ?u.7430
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit

nil
(forM [] fun x => act x) = List.forM [] act
m: Type Type ?u.7430
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
head✝: α
tail✝: List α
(forM (head :: tail) fun x => act x) = List.forM (head :: tail) act
m: Type Type ?u.7430
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit

nil
(forM [] fun x => act x) = List.forM [] act
m: Type Type ?u.7430
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
head✝: α
tail✝: List α
(forM (head :: tail) fun x => act x) = List.forM (head :: tail) act

Goals accomplished! 🐙
def
ex2: {β : Type} → (β → α → m β) → β → List α → m β
ex2
(
f: β → α → m β
f
:
β: Type
β
α: Type
α
m: Type → Type u_1
m
β: Type
β
) (
init: β
init
:
β: Type
β
) (
xs: List α
xs
:
List: Type → Type
List
α: Type
α
) :
m: Type → Type u_1
m
β: Type
β
:= do' let mut
y: β
y
:=
init: β
init
; for
x: α
x
in
xs: List α
xs
do' { y
f: β → α → m β
f
y: β
y
x: α
x
}; return
y: β
y
example: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (xs : List α) {β : Type} {init : β} [inst_1 : LawfulMonad m] (f : β → α → m β), ex2 f init xs = List.foldlM f init xs
example
[
LawfulMonad: (m : Type ?u.8985 → Type ?u.8984) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
f: β → α → m β
f
:
β: Type
β
α: Type
α
m: Type → Type u_1
m
β: Type
β
) :
ex2: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → {β : Type} → (β → α → m β) → β → List α → m β
ex2
f: β → α → m β
f
init: β
init
xs: List α
xs
=
xs: List α
xs
.
foldlM: {m : Type → Type u_1} → [inst : Monad m] → {s α : Type} → (s → α → m s) → s → List α → m s
foldlM
f: β → α → m β
f
init: β
init
:=
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
init: β
f: β α m β

ex2 f init xs = List.foldlM f init xs
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
init: β
f: β α m β

ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y) = List.foldlM f init xs
;
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
f: β α m β
init: β

nil
ExceptCpsT.runCatch (let y := init; StateT.run' (do forM [] fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y) = List.foldlM f init []
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
f: β α m β
head✝: α
tail✝: List α
init: β
ExceptCpsT.runCatch (let y := init; StateT.run' (do forM (head :: tail) fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y) = List.foldlM f init (head :: tail)
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
f: β α m β
init: β

nil
ExceptCpsT.runCatch (let y := init; StateT.run' (do forM [] fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y) = List.foldlM f init []
m: Type Type ?u.8948
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
β: Type
f: β α m β
head✝: α
tail✝: List α
init: β
ExceptCpsT.runCatch (let y := init; StateT.run' (do forM (head :: tail) fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y) = List.foldlM f init (head :: tail)

Goals accomplished! 🐙
@[simp] theorem
List.find?_cons: ∀ {x : α} {p : α → Bool} {xs : List α}, find? p (x :: xs) = if p x = true then some x else find? p xs
List.find?_cons
{
xs: List α
xs
:
List: Type ?u.17785 → Type ?u.17785
List
α: Type
α
} : (
x: α
x
::
xs: List α
xs
).
find?: {α : Type} → (α → Bool) → List α → Option α
find?
p: α → Bool
p
= if
p: α → Bool
p
x: α
x
then
some: {α : Type} → α → Option α
some
x: α
x
else
xs: List α
xs
.
find?: {α : Type} → (α → Bool) → List α → Option α
find?
p: α → Bool
p
:=
m: Type Type ?u.17767
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
x: α
p: α Bool
xs: List α

find? p (x :: xs) = if p x = true then some x else find? p xs
m: Type Type ?u.17767
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
x: α
p: α Bool
xs: List α
h: p x = false

false
find? p (x :: xs) = if false = true then some x else find? p xs
m: Type Type ?u.17767
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
x: α
p: α Bool
xs: List α
h: p x = true
find? p (x :: xs) = if true = true then some x else find? p xs
m: Type Type ?u.17767
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
x: α
p: α Bool
xs: List α
h: p x = false

false
find? p (x :: xs) = if false = true then some x else find? p xs
m: Type Type ?u.17767
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
x: α
p: α Bool
xs: List α
h: p x = true
find? p (x :: xs) = if true = true then some x else find? p xs

Goals accomplished! 🐙
example: ∀ {α : Type} (xs : List α) (p : α → Bool), Id.run (ExceptCpsT.runCatch do forM xs fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p xs
example
(
p: α → Bool
p
:
α: Type
α
Bool: Type
Bool
) :
Id.run: {α : Type} → Id α → α
Id.run
(do' for
x: α
x
in
xs: List α
xs
do' { if
p: α → Bool
p
x: α
x
then { return
some: {α : Type} → α → Option α
some
x: α
x
} };
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
none: {α : Type} → Option α
none
) =
xs: List α
xs
.
find?: {α : Type} → (α → Bool) → List α → Option α
find?
p: α → Bool
p
:=
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool

Id.run (ExceptCpsT.runCatch do forM xs fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p xs
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool

Id.run (ExceptCpsT.runCatch do forM xs fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p xs
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool

nil
Id.run (ExceptCpsT.runCatch do forM [] fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p []

Goals accomplished! 🐙
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool
x: α
tail✝: List α

cons
Id.run (ExceptCpsT.runCatch do forM (x :: tail) fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p (x :: tail)
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool
x: α
tail✝: List α
h: p x = false

cons.false
Id.run (ExceptCpsT.runCatch do forM (x :: tail) fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p (x :: tail)
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool
x: α
tail✝: List α
h: p x = true
Id.run (ExceptCpsT.runCatch do forM (x :: tail) fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p (x :: tail)
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool
x: α
tail✝: List α
h: p x = false

cons.false
Id.run (ExceptCpsT.runCatch do forM (x :: tail) fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p (x :: tail)
m: Type Type ?u.18571
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α Bool
x: α
tail✝: List α
h: p x = true
Id.run (ExceptCpsT.runCatch do forM (x :: tail) fun x => if p x = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.find? p (x :: tail)

Goals accomplished! 🐙
variable (
p: α → m Bool
p
:
α: Type
α
m: Type → Type ?u.38841
m
Bool: Type
Bool
) theorem
byCases_Bool_bind: ∀ {m : Type → Type u_1} [inst : Monad m] {β : Type} (x : m Bool) (f g : Bool → m β), f true = g true → f false = g false → x >>= f = x >>= g
byCases_Bool_bind
(
x: m Bool
x
:
m: Type → Type u_1
m
Bool: Type
Bool
) (
f: Bool → m β
f
g: Bool → m β
g
:
Bool: Type
Bool
m: Type → Type u_1
m
β: Type
β
) (
isTrue: f true = g true
isTrue
:
f: Bool → m β
f
true: Bool
true
=
g: Bool → m β
g
true: Bool
true
) (
isFalse: f false = g false
isFalse
:
f: Bool → m β
f
false: Bool
false
=
g: Bool → m β
g
false: Bool
false
) : (
x: m Bool
x
>>=
f: Bool → m β
f
) = (
x: m Bool
x
>>=
g: Bool → m β
g
) :=
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false

x >>= f = x >>= g
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false

x >>= f = x >>= g
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false

f = g
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b✝: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false
b: Bool

h
f b = g b
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b✝: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false
b: Bool

h
f b = g b
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false

h.true
f true = g true

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false

h.false
f false = g false

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false
this: f = g

x >>= f = x >>= g
m: Type Type u_1
α: Type
inst✝: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
β: Type
x: m Bool
f, g: Bool m β
isTrue: f true = g true
isFalse: f false = g false
this: f = g

x >>= g = x >>= g

Goals accomplished! 🐙
theorem
eq_findM: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (xs : List α) (p : α → m Bool) [inst_1 : LawfulMonad m], (ExceptCpsT.runCatch do forM xs fun x => do let b ← ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs
eq_findM
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] : (do' for
x: α
x
in
xs: List α
xs
do' { let
b: Bool
b
p: α → m Bool
p
x: α
x
; if
b: Bool
b
then { return
some: {α : Type} → α → Option α
some
x: α
x
} };
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
none: {α : Type} → Option α
none
) =
xs: List α
xs
.
findM?: {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → (α → m Bool) → List α → m (Option α)
findM?
p: α → m Bool
p
:=
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

nil
(ExceptCpsT.runCatch do forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p []

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p (x :: xs)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p (x :: xs)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a p x match a with | true => pure (some x) | false => List.findM? p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a p x match a with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a p x match a with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
;
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a p x match a with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = match true with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs
(ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = match false with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs

cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = match true with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findM? p xs
(ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = match false with | true => pure (some x) | false => ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

Goals accomplished! 🐙
def
ex3: [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3
[
Monad: (Type → Type u_1) → Type (max 1 u_1)
Monad
m: Type → Type u_1
m
] (
p: α → m Bool
p
:
α: Type
α
m: Type → Type u_1
m
Bool: Type
Bool
) (
xss: List (List α)
xss
:
List: Type → Type
List
(
List: Type → Type
List
α: Type
α
)) :
m: Type → Type u_1
m
(
Option: Type → Type
Option
α: Type
α
) := do' for
xs: List α
xs
in
xss: List (List α)
xss
do' { for
x: α
x
in
xs: List α
xs
do' { let
b: Bool
b
p: α → m Bool
p
x: α
x
; if
b: Bool
b
then { return
some: {α : Type} → α → Option α
some
x: α
x
} } };
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
none: {α : Type} → Option α
none
theorem
eq_findSomeM_findM: ∀ {m : Type → Type u_1} {α : Type} [inst : Monad m] (p : α → m Bool) [inst_1 : LawfulMonad m] (xss : List (List α)), ex3 p xss = List.findSomeM? (fun xs => List.findM? p xs) xss
eq_findSomeM_findM
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
xss: List (List α)
xss
:
List: Type → Type
List
(
List: Type → Type
List
α: Type
α
)) :
ex3: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3
p: α → m Bool
p
xss: List (List α)
xss
=
xss: List (List α)
xss
.
findSomeM?: {m : Type → Type u_1} → [inst : Monad m] → {α β : Type} → (α → m (Option β)) → List α → m (Option β)
findSomeM?
(fun
xs: List α
xs
=>
xs: List α
xs
.
findM?: {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → (α → m Bool) → List α → m (Option α)
findM?
p: α → m Bool
p
) :=
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
xss: List (List α)

ex3 p xss = List.findSomeM? (fun xs => List.findM? p xs) xss
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
xss: List (List α)

(ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
xss: List (List α)

(ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

nil
(ExceptCpsT.runCatch do forM [] fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) []

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM (xs :: xss) fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) (xs :: xss)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a List.findM? p xs match a with | some b => pure (some b) | none => List.findSomeM? (fun xs => List.findM? p xs) xss
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a List.findM? p xs match a with | some b => pure (some b) | none => List.findSomeM? (fun xs => List.findM? p xs) xss
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a List.findM? p xs match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xs: List α
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons
(ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool
xss: List (List α)
ih: (ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = List.findSomeM? (fun xs => List.findM? p xs) xss

cons.nil
(ExceptCpsT.runCatch do forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.cons
(ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let x_1 p x let a ExceptCpsT.runCatch do if x_1 = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
;
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
(ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do if true = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
xss: List (List α)
x: α
xs: List α
ih: (ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
(ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)) = do let a ExceptCpsT.runCatch do if false = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none) match a with | some b => pure (some b) | none => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

Goals accomplished! 🐙
def
List.untilM: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → (α → m Bool) → List α → m Unit
List.untilM
(
p: α → m Bool
p
:
α: Type
α
m: Type → Type u_1
m
Bool: Type
Bool
) :
List: Type → Type
List
α: Type
α
m: Type → Type u_1
m
Unit: Type
Unit
| [] =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
|
a: α
a
::
as: List α
as
=>
p: α → m Bool
p
a: α
a
>>= fun |
true: Bool
true
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
|
false: Bool
false
=>
as: List α
as
.
untilM: (α → m Bool) → List α → m Unit
untilM
p: α → m Bool
p
theorem
eq_untilM: ∀ [inst : LawfulMonad m], ExceptCpsT.runCatch (forM xs fun x => do let b ← ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs
eq_untilM
[
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] : (do' for
x: α
x
in
xs: List α
xs
do' { let
b: Bool
b
p: α → m Bool
p
x: α
x
; if
b: Bool
b
then { break } }) =
xs: List α
xs
.
untilM: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → (α → m Bool) → List α → m Unit
untilM
p: α → m Bool
p
:=
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs: List α
act: α m Unit
p: α m Bool

nil
ExceptCpsT.runCatch (forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p []

Goals accomplished! 🐙
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs

cons
ExceptCpsT.runCatch (forM (x :: xs) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p (x :: xs)
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = do let x p x match x with | true => pure () | false => List.untilM p xs
;
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = do let x p x match x with | true => pure () | false => List.untilM p xs
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = do let x p x match x with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
ih: ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = List.untilM p xs

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = do let x p x match x with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
;
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α

cons
(do let a p x ExceptCpsT.runCatch do if a = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = do let x p x match x with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α

cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = match true with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
(ExceptCpsT.runCatch do if false = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = match false with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α

cons.isTrue
(ExceptCpsT.runCatch do if true = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = match true with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))
m: Type Type u_1
α: Type
inst✝¹: Monad m
ma, ma': m α
b: Bool
xs✝: List α
act: α m Unit
p: α m Bool
x: α
xs: List α
(ExceptCpsT.runCatch do if false = true then throw () else ExceptCpsT.lift (pure ()) forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ())) = match false with | true => pure () | false => ExceptCpsT.runCatch (forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw () else ExceptCpsT.lift (pure ()))

Goals accomplished! 🐙
/- The notation `[0:10]` is a range from 0 to 10 (exclusively). -/
0 0 1 0 2 0 3 0 4 0 5
do' for
x: Nat
x
in [
0: Nat
0
:
10: Nat
10
] do' { if
x: Nat
x
>
5: Nat
5
then { break }; for
y: Nat
y
in [
0: Nat
0
:
x: Nat
x
] do' {
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
y: Nat
y
; break };
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
x: Nat
x
}
0 1 2 3 4 5
do' for
x: Nat
x
in [
0: Nat
0
:
10: Nat
10
] do' { if
x: Nat
x
>
5: Nat
5
then { break };
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
x: Nat
x
}
1 3 5
do' for
x: Nat
x
in [
0: Nat
0
:
10: Nat
10
] do' { if
x: Nat
x
%
2: Nat
2
==
0: Nat
0
then { continue }; if
x: Nat
x
>
5: Nat
5
then { break };
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
x: Nat
x
}
1 3 5
do' for
x: Nat
x
in [
0: Nat
0
:
10: Nat
10
] do' { if
x: Nat
x
%
2: Nat
2
==
0: Nat
0
then { continue }; if
x: Nat
x
>
5: Nat
5
then { return () };
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
x: Nat
x
} -- set_option trace.compiler.ir.init true def
ex1: List Nat → Nat → Id Nat
ex1
(
xs: List Nat
xs
:
List: Type → Type
List
Nat: Type
Nat
) (
z: Nat
z
:
Nat: Type
Nat
) :
Id: Type → Type
Id
Nat: Type
Nat
:= do' let mut
s1: Nat
s1
:=
0: Nat
0
; let mut
s2: Nat
s2
:=
0: Nat
0
; for
x: Nat
x
in
xs: List Nat
xs
do' { if
x: Nat
x
%
2: Nat
2
==
0: Nat
0
then { continue }; if
x: Nat
x
==
z: Nat
z
then { return
s1: Nat
s1
}; s1 :=
s1: Nat
s1
+
x: Nat
x
; s2 :=
s2: Nat
s2
+
s1: Nat
s1
}; return (
s1: Nat
s1
+
s2: Nat
s2
) /- Adding `repeat` and `while` statements -/ -- The "partial" keyword allows users to define non-terminating functions in Lean. -- However, we currently cannot reason about them. @[specialize] partial def
loopForever: [inst : Monad m] → (Unit → m Unit) → m Unit
loopForever
[
Monad: (Type → Type u_1) → Type (max 1 u_1)
Monad
m: Type → Type u_1
m
] (
f: Unit → m Unit
f
:
Unit: Type
Unit
m: Type → Type u_1
m
Unit: Type
Unit
) :
m: Type → Type u_1
m
Unit: Type
Unit
:=
f: Unit → m Unit
f
(): Unit
()
*>
loopForever: [inst : Monad m] → (Unit → m Unit) → m Unit
loopForever
f: Unit → m Unit
f
-- `Loop'` is a "helper" type. It is similar to `Unit`. -- Its `ForM` instance produces an "infinite" sequence of units. inductive
Loop': Type
Loop'
where |
mk: Loop'
mk
:
Loop': Type
Loop'
instance: {m : Type → Type u_1} → ForM m Loop' Unit
instance
:
ForM: (Type → Type u_1) → Type → outParam Type → Type (max (max 1 u_1) 0)
ForM
m: Type → Type u_1
m
Loop': Type
Loop'
Unit: Type
Unit
where forM _
f: Unit → m PUnit
f
:=
loopForever: {m : Type → Type u_1} → [inst : Monad m] → (Unit → m Unit) → m Unit
loopForever
f: Unit → m PUnit
f
macro:0 "repeat"
s: TSyntax `stmt
s
:stmt:1 : stmt => `(stmt| for u in Loop'.mk do' $
s: TSyntax `stmt
s
)
0 1 2 3 4 5 6 7 8 9 10 11
do' let mut
i: Nat
i
:=
0: Nat
0
; repeat { if
i: Nat
i
>
10: Nat
10
then { break };
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
i: Nat
i
; i :=
i: Nat
i
+
1: Nat
1
}; return
i: Nat
i
macro:0 "while"
c: TSyntax `term
c
:term "do'"
s: TSyntax `stmt
s
:stmt:1 : stmt => `(stmt| repeat { unless $
c: TSyntax `term
c
do' break; { $
s: TSyntax `stmt
s
} })
0 1 2 3 4 5 6 7 8 9 10
do' let mut
i: Nat
i
:=
0: Nat
0
; while (
i: Nat
i
<
10: Nat
10
) do' {
IO.println: {α : Type} → [inst : ToString α] → α → IO Unit
IO.println
i: Nat
i
; i :=
i: Nat
i
+
1: Nat
1
}; return
i: Nat
i
import Do.For  -- import `runCatch`
import Lean
import Aesop

Formalization of Extended do Translation

An intrinsically typed representation of the paper's syntax of do statements as well of their translation functions and an equivalence proof thereof to a simple denotational semantics.

Contexts

We represent contexts as lists of types and assignments of them as heterogeneous lists over these types. As is common with lists, contexts grow to the left in our presentation. The following encoding of heterogeneous lists avoids the universe bump of the usual inductive definition (https://lists.chalmers.se/pipermail/agda/2010/001826.html).

def 
HList: List (Type u) → Type u
HList
:
List: Type (u + 1) → Type (u + 1)
List
(
Type u: Type (u + 1)
Type u
)
Type u: Type (u + 1)
Type u
| [] =>
PUnit: Type u
PUnit
|
α: Type u
α
::
αs: List (Type u)
αs
=>
α: Type u
α
×
HList: List (Type u) → Type u
HList
αs: List (Type u)
αs
@[matchPattern] abbrev
HList.nil: HList []
HList.nil
:
HList: List (Type u_1) → Type u_1
HList
[]: List (Type u_1)
[]
:=
⟨⟩: PUnit
⟨⟩
@[matchPattern] abbrev
HList.cons: {α : Type u_1} → {αs : List (Type u_1)} → α → HList αs → HList (α :: αs)
HList.cons
(
a: α
a
:
α: Type u_1
α
) (
as: HList αs
as
:
HList: List (Type u_1) → Type u_1
HList
αs: List (Type u_1)
αs
) :
HList: List (Type u_1) → Type u_1
HList
(
α: Type u_1
α
::
αs: List (Type u_1)
αs
) := (
a: α
a
,
as: HList αs
as
)

We overload the common list notations :: and [e, ...] for HList.

infixr:67 " :: " => HList.cons

syntax (name := hlistCons) "[" term,* "]" : term
macro_rules (kind := hlistCons)
  | `([])          => 
`(HList.nil): Lean.MacroM Lean.Syntax
`(HList.nil)
| `([$
x: Lean.TSyntax `term
x
]) => `(HList.cons $
x: Lean.TSyntax `term
x
[]) | `([$
x: Lean.TSyntax `term
x
, $
xs: Lean.Syntax.TSepArray `term ","
xs
,*]) => `(HList.cons $
x: Lean.TSyntax `term
x
[$
xs: Lean.Syntax.TSepArray `term ","
xs
,*])

Lean's very general, heterogeneous definition of ++ causes some issues with our overloading above in terms such as a ++ [b], so we restrict it to the List interpretation in the following, which is sufficient for our purposes.

local macro_rules
  | `($
a: Lean.TSyntax `term
a
++ $
b: Lean.TSyntax `term
b
) => `(List.append $
a: Lean.TSyntax `term
a
$
b: Lean.TSyntax `term
b
) abbrev
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
:=
HList: List (Type u_1) → Type u_1
HList
Γ: List (Type u_1)
Γ

Updating a heterogeneous list at a given, guaranteed in-bounds index.

def 
HList.set: {αs : List (Type u)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
HList.set
: {
αs: List (Type u)
αs
:
List: Type (u + 1) → Type (u + 1)
List
(
Type u: Type (u + 1)
Type u
)}
HList: List (Type u) → Type u
HList
αs: List (Type u)
αs
(
i: Fin (List.length αs)
i
:
Fin: Nat → Type
Fin
αs: List (Type u)
αs
.
length: {α : Type (u + 1)} → List α → Nat
length
)
αs: List (Type u)
αs
.
get: {α : Type (u + 1)} → (as : List α) → Fin (List.length as) → α
get
i: Fin (List.length αs)
i
HList: List (Type u) → Type u
HList
αs: List (Type u)
αs
| _ :: _, _ ::
as: HList a✝
as
, ⟨
0: Nat
0
, _⟩,
b: List.get (α✝ :: a✝) { val := 0, isLt := isLt✝ }
b
=>
b: List.get (α✝ :: a✝) { val := 0, isLt := isLt✝ }
b
::
as: HList a✝
as
| _ :: _,
a: α✝
a
::
as: HList a✝
as
, ⟨
Nat.succ: Nat → Nat
Nat.succ
n: Nat
n
,
h: Nat.succ n < List.length (α✝ :: a✝)
h
⟩,
b: List.get (α✝ :: a✝) { val := Nat.succ n, isLt := h }
b
=>
a: α✝
a
::
set: {αs : List (Type u)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
as: HList a✝
as
n: Nat
n
,
Nat.le_of_succ_le_succ: ∀ {n m : Nat}, Nat.succ n ≤ Nat.succ m → n ≤ m
Nat.le_of_succ_le_succ
h: Nat.succ n < List.length (α✝ :: a✝)
h
b: List.get (α✝ :: a✝) { val := Nat.succ n, isLt := h }
b
| [], [], _, _ =>
[]: HList []
[]

We write for empty contexts and assignments and Γ ⊢ α for the type of values of type α under the context Γ

  • that is, the function type from an assignment to α.
instance: EmptyCollection (Assg ∅)
instance
:
EmptyCollection: Type u_1 → Type u_1
EmptyCollection
(
Assg: List (Type u_1) → Type u_1
Assg
: List (Type u_1)
) where emptyCollection :=
[]: HList []
[]
instance: CoeSort (List (Type u)) (Type u)
instance
:
CoeSort: Type (u + 1) → outParam (Type (u + 1)) → Type (u + 1)
CoeSort
(
List: Type (u + 1) → Type (u + 1)
List
(
Type u: Type (u + 1)
Type u
)) (
Type u: Type (u + 1)
Type u
) := ⟨
Assg: List (Type u) → Type u
Assg
notation:30
Γ: Lean.TSyntax `term
Γ
" ⊢ "
α: Lean.TSyntax `term
α
=> Assg
Γ: Lean.TSyntax `term
Γ
α: Lean.TSyntax `term
α
def
Assg.drop: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (α :: Γ) → Assg Γ
Assg.drop
:
Assg: List (Type u_1) → Type u_1
Assg
(
α: Type u_1
α
::
Γ: List (Type u_1)
Γ
)
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
| _ ::
as: HList Γ
as
=>
as: HList Γ
as

In one special case, we will need to manipulate contexts from the right, i.e. the outermost scope.

def 
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
(
x: α
x
:
α: Type u_1
α
) : {
Γ: List (Type u_1)
Γ
:
_: Type (u_1 + 1)
_
}
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
Assg: List (Type u_1) → Type u_1
Assg
(
Γ: List (Type u_1)
Γ
++ [
α: Type u_1
α
]) | [], [] => [
x: α
x
] | _ :: _,
a: α✝
a
::
as: HList a✝
as
=>
a: α✝
a
::
extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
extendBot
x: α
x
as: HList a✝
as
def
Assg.dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
: {
Γ: List (Type u_1)
Γ
:
_: Type (u_1 + 1)
_
}
Assg: List (Type u_1) → Type u_1
Assg
(
Γ: List (Type u_1)
Γ
++ [
α: Type u_1
α
])
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
| [], _ =>
[]: HList []
[]
| _ ::
_: List (Type u_1)
_
,
a: α✝
a
::
as: HList (List.append tail✝ [α])
as
=>
a: α✝
a
::
dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
dropBot
as: HList (List.append tail✝ [α])
as

Intrinsically Typed Representation of do Statements

where

  • m: base monad
  • ω: return type, m ω is the type of the entire do block
  • Γ: do-local immutable context
  • Δ: do-local mutable context
  • b: break allowed
  • c: continue allowed
  • α: local result type, m α is the type of the statement

The constructor signatures are best understood by comparing them with the corresponding typing rules in the paper. Note that the choice of de Bruijn indices changes/simplifies some parts, such as obviating freshness checks (x ∉ Δ).

inductive 
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
(
m: Type → Type u
m
:
Type: Type 1
Type
Type u: Type (u + 1)
Type u
) (
ω: Type
ω
:
Type: Type 1
Type
) : (
Γ: List Type
Γ
Δ: List Type
Δ
:
List: Type 1 → Type 1
List
Type: Type 1
Type
) (
b: Bool
b
c: Bool
c
:
Bool: Type
Bool
) (
α: Type
α
:
Type: Type 1
Type
)
Type _: Type ((max 1 u) + 1)
Type _
where |
expr: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
expr
(
e: Assg Γ → Assg Δ → m α
e
:
Γ: List Type
Γ
Δ: List Type
Δ
m: Type → Type u
m
α: Type
α
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
|
bind: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
bind
(
s: Stmt m ω Γ Δ b c α
s
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
) (
s': Stmt m ω (α :: Γ) Δ b c β
s'
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
(
α: Type
α
::
Γ: List Type
Γ
)
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
-- let _ ← s; s' |
letmut: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
letmut
(
e: Assg Γ → Assg Δ → α
e
:
Γ: List Type
Γ
Δ: List Type
Δ
α: Type
α
) (
s: Stmt m ω Γ (α :: Δ) b c β
s
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
(
α: Type
α
::
Δ: List Type
Δ
)
b: Bool
b
c: Bool
c
β: Type
β
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
-- let mut _ := e; s |
assg: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
assg
(
x: failed to pretty print expression (use 'set_option pp.rawOnError true' for raw representation)
x
:
Fin: Nat → Type
Fin
Δ: List Type
Δ
Δ.length: Nat
.
length: {α : Type 1} → List α → Nat
length
) (
e: failed to pretty print expression (use 'set_option pp.rawOnError true' for raw representation)
e
:
Γ: List Type
Γ
Δ: List Type
Δ
Δ: List Type
Δ
.
get: {α : Type 1} → (as : List α) → Fin (List.length as) → α
get
x: failed to pretty print expression (use 'set_option pp.rawOnError true' for raw representation)
x
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
Unit: Type
Unit
-- x := e |
ite: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
ite
(
e: Assg Γ → Assg Δ → Bool
e
:
Γ: List Type
Γ
Δ: List Type
Δ
Bool: Type
Bool
) (
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
-- if e then s₁ else s₂ |
ret: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
ret
(
e: Assg Γ → Assg Δ → ω
e
:
Γ: List Type
Γ
Δ: List Type
Δ
ω: Type
ω
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
-- return e |
sfor: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
sfor
(
e: Assg Γ → Assg Δ → List α
e
:
Γ: List Type
Γ
Δ: List Type
Δ
List: Type → Type
List
α: Type
α
) (
s: Stmt m ω (α :: Γ) Δ true true Unit
s
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
(
α: Type
α
::
Γ: List Type
Γ
)
Δ: List Type
Δ
true: Bool
true
true: Bool
true
Unit: Type
Unit
) :
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
Unit: Type
Unit
-- for _ in e do s |
sbreak: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
sbreak
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
true: Bool
true
c: Bool
c
α: Type
α
-- break |
scont: {m : Type → Type u} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
scont
:
Stmt: (Type → Type u) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u)
Stmt
m: Type → Type u
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
true: Bool
true
α: Type
α
-- continue

Neutral statements are a restriction of the above type.

inductive 
Neut: Type → Type → Bool → Bool → Type
Neut
(
ω: Type
ω
α: Type
α
:
Type: Type 1
Type
) : (
b: Bool
b
c: Bool
c
:
Bool: Type
Bool
)
Type _: Type 1
Type _
where |
val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
val
(
a: α
a
:
α: Type
α
) :
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
c: Bool
c
|
ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
ret
(
o: ω
o
:
ω: Type
ω
) :
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
c: Bool
c
|
rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
rbreak
:
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
true: Bool
true
c: Bool
c
|
rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
rcont
:
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
true: Bool
true

We elide Neut.val where unambiguous.

instance: {α ω : Type} → {b c : Bool} → Coe α (Neut ω α b c)
instance
:
Coe: Type → Type → Type
Coe
α: Type
α
(
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
c: Bool
c
) := ⟨
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
instance: {α ω : Type} → {b c : Bool} → Coe (Id α) (Neut ω α b c)
instance
:
Coe: Type → Type → Type
Coe
(
Id: Type → Type
Id
α: Type
α
) (
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
c: Bool
c
) := ⟨
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val

We write e[ρ][σ] for the substitution of both contexts in e, a simple application in this encoding. σ[x ↦ v] updates σ at v (a de Bruijn index).

macro:max (priority := high) 
e: Lean.TSyntax `term
e
:term:max noWs "["
ρ: Lean.TSyntax `term
ρ
:term "]" "["
σ: Lean.TSyntax `term
σ
:term "]" : term => `($
e: Lean.TSyntax `term
e
$
ρ: Lean.TSyntax `term
ρ
$
σ: Lean.TSyntax `term
σ
) macro:max (priority := high)
σ: Lean.TSyntax `term
σ
:term:max noWs "["
x: Lean.TSyntax `term
x
:term " ↦ "
v: Lean.TSyntax `term
v
:term "]" : term => `(HList.set $
σ: Lean.TSyntax `term
σ
$
x: Lean.TSyntax `term
x
$
v: Lean.TSyntax `term
v
)

Dynamic Evaluation Function

A direct encoding of the paper's operational semantics as a denotational function, generalized over an arbitrary monad. Note that the immutable context ρ is accumulated (v :: ρ) and passed explicitly instead of immutable bindings being substituted immediately as that is a better match for the above definition of Stmt. Iteration over the values of the given list in the for case introduces a nested, mutually recursive helper function, with termination of the mutual bundle following from a size argument over the statement primarily and the length of the list in the for case secondarily.

@[simp] def 
Stmt.eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
Stmt.eval
[
Monad: (Type ?u.21808 → Type ?u.21807) → Type (max (?u.21808 + 1) ?u.21807)
Monad
m: Type → Type u_1
m
] (
ρ: Assg Γ
ρ
:
Assg: List Type → Type
Assg
Γ: List Type
Γ
) :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
Assg: List Type → Type
Assg
Δ: List Type
Δ
m: Type → Type u_1
m
(
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
b: Bool
b
c: Bool
c
×
Assg: List Type → Type
Assg
Δ: List Type
Δ
) |
expr: {m : Type → Type ?u.32858} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
expr
e: Assg Γ → Assg Δ → m α
e
,
σ: Assg Δ
σ
=>
e: Assg Γ → Assg Δ → m α
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
] >>= fun
v: α
v
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
v: α
v
,
σ: Assg Δ
σ
|
bind: {m : Type → Type ?u.33084} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
bind
s: Stmt m ω Γ Δ b c α✝
s
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
,
σ: Assg Δ
σ
=> -- defining this part as a separate definition helps Lean with the termination proof let rec @[simp]
cont: (α✝ → Assg Δ → m (Neut ω α b c × Assg Δ)) → Neut ω α✝ b c × Assg Δ → m (Neut ω α b c × Assg Δ)
cont
val: α✝ → Assg Δ → m (Neut ω α b c × Assg Δ)
val
|
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
v: α✝
v
,
σ': Assg Δ
σ'
=>
val: α✝ → Assg Δ → m (Neut ω α b c × Assg Δ)
val
v: α✝
v
σ': Assg Δ
σ'
-- the `Neut` type family forces us to repeat these cases as the LHS/RHS indices are not identical |
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ': Assg Δ
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ': Assg Δ
σ'
|
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ': Assg Δ
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ': Assg Δ
σ'
|
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ': Assg Δ
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ': Assg Δ
σ'
s: Stmt m ω Γ Δ b c α✝
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
>>=
cont: (α✝ → Assg Δ → m (Neut ω α b c × Assg Δ)) → Neut ω α✝ b c × Assg Δ → m (Neut ω α b c × Assg Δ)
cont
(fun
v: α✝
v
σ': Assg Δ
σ'
=>
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
(
v: α✝
v
::
ρ: Assg Γ
ρ
)
σ': Assg Δ
σ'
) |
letmut: {m : Type → Type ?u.33463} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
,
σ: Assg Δ
σ
=>
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
(
e: Assg Γ → Assg Δ → α✝
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
],
σ: Assg Δ
σ
) >>= fun
r: Neut ω α b c
r
,
σ': Assg (α✝ :: Δ)
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
r: Neut ω α b c
r
,
σ': Assg (α✝ :: Δ)
σ'
.
drop: {α : Type} → {Γ : List Type} → Assg (α :: Γ) → Assg Γ
drop
-- `x` is a valid de Bruijn index into `σ` by definition of `assg` |
assg: {m : Type → Type ?u.35837} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
,
σ: Assg Δ
σ
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
,
σ: Assg Δ
σ
[
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
]]⟩ |
ite: {m : Type → Type ?u.36112} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
,
σ: Assg Δ
σ
=> if
e: Assg Γ → Assg Δ → Bool
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
] then
s₁: Stmt m ω Γ Δ b c α
s₁
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
else
s₂: Stmt m ω Γ Δ b c α
s₂
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
|
ret: {m : Type → Type ?u.36353} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
ret
e: Assg Γ → Assg Δ → ω
e
,
σ: Assg Δ
σ
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
e: Assg Γ → Assg Δ → ω
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
],
σ: Assg Δ
σ
|
sfor: {m : Type → Type ?u.36472} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
,
σ: Assg Δ
σ
=> let rec @[simp]
go: Assg Δ → List α✝ → m (Neut ω Unit b c × Assg Δ)
go
σ: Assg Δ
σ
| [] =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
,
σ: Assg Δ
σ
|
a: α✝
a
::
as: List α✝
as
=>
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
(
a: α✝
a
::
ρ: Assg Γ
ρ
)
σ: Assg Δ
σ
>>= fun | ⟨(),
σ': Assg Δ
σ'
=>
go: Assg Δ → List α✝ → m (Neut ω Unit b c × Assg Δ)
go
σ': Assg Δ
σ'
as: List α✝
as
|
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ': Assg Δ
σ'
=>
go: Assg Δ → List α✝ → m (Neut ω Unit b c × Assg Δ)
go
σ': Assg Δ
σ'
as: List α✝
as
|
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ': Assg Δ
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
,
σ': Assg Δ
σ'
|
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ': Assg Δ
σ'
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ': Assg Δ
σ'
go: Assg Δ → List α✝ → m (Neut ω Unit b c × Assg Δ)
go
σ: Assg Δ
σ
e: Assg Γ → Assg Δ → List α✝
e
[
ρ: Assg Γ
ρ
][
σ: Assg Δ
σ
] |
sbreak: {m : Type → Type ?u.36610} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
sbreak
,
σ: Assg Δ
σ
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ: Assg Δ
σ
|
scont: {m : Type → Type ?u.36721} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
scont
,
σ: Assg Δ
σ
=>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ: Assg Δ
σ
termination_by eval s _ => (
sizeOf: {α : Type (max 1 u_1)} → [self : SizeOf α] → α → Nat
sizeOf
s: Stmt m ω Γ Δ b c α
s
,
0: Nat
0
) eval.go as => (
sizeOf: {α : Type (max 1 u_1)} → [self : SizeOf α] → α → Nat
sizeOf
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
,
as: List α✝
as
.
length: {α : Type} → List α → Nat
length
)

At the top-level statement, the contexts are empty, no loop control flow statements are allowed, and the return and result type are identical.

abbrev 
Do: (Type → Type u_1) → Type → Type (max 1 u_1)
Do
m: Type → Type u_1
m
α: Type
α
:=
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
α: Type
α
: List Type
: List Type
false: Bool
false
false: Bool
false
α: Type
α
def
Do.eval: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → Do m α → m α
Do.eval
[
Monad: (Type → Type u_1) → Type (max 1 u_1)
Monad
m: Type → Type u_1
m
] (
s: Do m α
s
:
Do: (Type → Type u_1) → Type → Type (max 1 u_1)
Do
m: Type → Type u_1
m
α: Type
α
) :
m: Type → Type u_1
m
α: Type
α
:= -- corresponds to the reduction rule `do s ⇒ v` in the paper
Stmt.eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
Stmt.eval
: Assg ∅
s: Do m α
s
: Assg ∅
>>= fun |
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
, _⟩ =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
a: α
a
|
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: α
o
, _⟩ =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
o: α
o
notation "⟦"
s: Lean.TSyntax `term
s
"⟧" => Do.eval
s: Lean.TSyntax `term
s

Translation Functions

We adjust the immutable context where necessary. The mutable context does not have to be adjusted.

@[simp] def 
Stmt.mapAssg: (_x : (Γ' : List Type) ×' (Γ : List Type) ×' (m : Type → Type u_1) ×' (ω : Type) ×' (Δ : List Type) ×' (b : Bool) ×' (c : Bool) ×' (β : Type) ×' (_ : Assg Γ' → Assg Γ) ×' Stmt m ω Γ Δ b c β) → Stmt _x.2.2.1 _x.2.2.2.1 _x.1 _x.2.2.2.2.1 _x.2.2.2.2.2.1 _x.2.2.2.2.2.2.1 _x.2.2.2.2.2.2.2.1
Stmt.mapAssg
(
f: Assg Γ' → Assg Γ
f
:
Assg: List Type → Type
Assg
Γ': List Type
Γ'
Assg: List (Type ?u.111841) → Type ?u.111841
Assg
Γ: List Type
Γ
) :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ': List Type
Γ'
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
|
expr: {m : Type → Type ?u.116778} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
expr
e: Assg Γ → Assg Δ → m β
e
=>
expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
expr
(
e: Assg Γ → Assg Δ → m β
e
f: Assg Γ' → Assg Γ
f
) |
bind: {m : Type → Type ?u.116872} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
bind
s₁: Stmt m ω Γ Δ b c α✝
s₁
s₂: Stmt m ω (α✝ :: Γ) Δ b c β
s₂
=>
bind: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
bind
(
s₁: Stmt m ω Γ Δ b c α✝
s₁
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
f: Assg Γ' → Assg Γ
f
) (
s₂: Stmt m ω (α✝ :: Γ) Δ b c β
s₂
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
(fun (
a: α✝
a
::
as: HList Γ'
as
) => (
a: α✝
a
::
f: Assg Γ' → Assg Γ
f
as: HList Γ'
as
))) |
letmut: {m : Type → Type ?u.117466} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) b c β
s
=>
letmut: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
letmut
(
e: Assg Γ → Assg Δ → α✝
e
f: Assg Γ' → Assg Γ
f
) (
s: Stmt m ω Γ (α✝ :: Δ) b c β
s
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
f: Assg Γ' → Assg Γ
f
) |
assg: {m : Type → Type ?u.117619} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
=>
assg: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
assg
x: Fin (List.length Δ)
x
(
e: Assg Γ → Assg Δ → List.get Δ x
e
f: Assg Γ' → Assg Γ
f
) |
ite: {m : Type → Type ?u.117739} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ b c β
s₁
s₂: Stmt m ω Γ Δ b c β
s₂
=>
ite: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
ite
(
e: Assg Γ → Assg Δ → Bool
e
f: Assg Γ' → Assg Γ
f
) (
s₁: Stmt m ω Γ Δ b c β
s₁
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
f: Assg Γ' → Assg Γ
f
) (
s₂: Stmt m ω Γ Δ b c β
s₂
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
f: Assg Γ' → Assg Γ
f
) |
ret: {m : Type → Type ?u.117891} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
ret
e: Assg Γ → Assg Δ → ω
e
=>
ret: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
ret
(
e: Assg Γ → Assg Δ → ω
e
f: Assg Γ' → Assg Γ
f
) |
sfor: {m : Type → Type ?u.117980} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
sfor: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
sfor
(
e: Assg Γ → Assg Δ → List α✝
e
f: Assg Γ' → Assg Γ
f
) (
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
.
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
(fun (
a: α✝
a
::
as: HList Γ'
as
) => (
a: α✝
a
::
f: Assg Γ' → Assg Γ
f
as: HList Γ'
as
))) |
sbreak: {m : Type → Type ?u.118585} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
sbreak
=>
sbreak: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
sbreak
|
scont: {m : Type → Type ?u.118632} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
scont
=>
scont: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
scont

Let us write f ∘ₑ e for the composition of f : α → β with e : Γ ⊢ Δ ⊢ α, which we will use to rewrite embedded terms.

infixr:90 " ∘ₑ "  => fun f e => fun ρ σ => f e[ρ][σ]

The formalization of S creates some technical hurdles. Because it operates on the outer-most mutable binding, we have to operate on that context from the right, from which we lose some helpful definitional equalities and have to rewrite types using nested proofs instead.

The helper function shadowSnd is particularly interesting because it shows how the shadowing in translation rules (S2) and (S9) is expressed in our de Bruijn encoding: The context α :: β :: α :: Γ corresponds, in this order, to the y that has just been bound to the value of get, then x from the respective rule, followed by the y of the outer scope. We encode the shadowing of y by dropping the third element from the context as well as the assignment. We are in fact forced to do so because the corresponding branches of S would not otherwise typecheck. The only mistake we could still make is to drop the wrong α value from the assignment, which (speaking from experience) would eventually be caught by the correctness proof.

@[simp] def 
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
[
Monad: (Type ?u.178851 → Type ?u.178850) → Type (max (?u.178851 + 1) ?u.178850)
Monad
m: Type → Type
m
] :
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
m: Type → Type
m
ω: Type
ω
Γ: List Type
Γ
(
Δ: List Type
Δ
++ [
α: Type
α
])
b: Bool
b
c: Bool
c
β: Type
β
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
(
StateT: Type → (Type → Type) → Type → Type
StateT
α: Type
α
m: Type → Type
m
)
ω: Type
ω
(
α: Type
α
::
Γ: List Type
Γ
)
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
/-(S1)-/ |
Stmt.expr: {m : Type → Type ?u.189083} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg (List.append Δ [α]) → m β
e
=>
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
StateT.lift: {σ : Type} → {m : Type → Type} → [inst : Monad m] → {α : Type} → m α → StateT σ m α
StateT.lift
∘ₑ
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → m β
e
) /-(S2)-/ |
Stmt.bind: {m : Type → Type ?u.189320} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s₁: Stmt m ω Γ (List.append Δ [α]) b c α✝
s₁
s₂: Stmt m ω (α✝ :: Γ) (List.append Δ [α]) b c β
s₂
=>
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s₁: Stmt m ω Γ (List.append Δ [α]) b c α✝
s₁
) (
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
get: {σ : Type} → {m : Type → Type} → [self : MonadState σ m] → m σ
get
)) (
Stmt.mapAssg: {Γ' Γ : List Type} → {m : Type → Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
Stmt.mapAssg
shadowSnd: {β : Type} → Assg (α :: β :: α :: Γ) → Assg (α :: β :: Γ)
shadowSnd
(
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s₂: Stmt m ω (α✝ :: Γ) (List.append Δ [α]) b c β
s₂
))) /-(S3)-/ |
Stmt.letmut: {m : Type → Type ?u.189575} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg (List.append Δ [α]) → α✝
e
s: Stmt m ω Γ (α✝ :: List.append Δ [α]) b c β
s
=>
Stmt.letmut: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
(
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → α✝
e
) (
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s: Stmt m ω Γ (α✝ :: List.append Δ [α]) b c β
s
) |
Stmt.assg: {m : Type → Type ?u.189691} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length (List.append Δ [α]))
x
e: Assg Γ → Assg (List.append Δ [α]) → List.get (List.append Δ [α]) x
e
=> if
h: ¬x.val < List.length Δ
h
:
x: Fin (List.length (List.append Δ [α]))
x
<
Δ: List Type
Δ
.
length: {α : Type 1} → List α → Nat
length
then /-(S4)-/
Stmt.assg: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length (List.append Δ [α]))
x
,
h: x.val < List.length Δ
h
⟩ (fun (
y: α
y
::
ρ: HList Γ
ρ
)
σ: Assg Δ
σ
=>
List.get_append_left: ∀ {α : Type 1} {i : Nat} (as bs : List α) (h : i < List.length as) {h' : i < List.length (as ++ bs)}, List.get (as ++ bs) { val := i, isLt := h' } = List.get as { val := i, isLt := h }
List.get_append_left
..
e: Assg Γ → Assg (List.append Δ [α]) → List.get (List.append Δ [α]) x
e
ρ: HList Γ
ρ
(
Assg.extendBot: {α : Type} → α → {Γ : List Type} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
y: α
y
σ: Assg Δ
σ
)) else /-(S5)-/
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
set: {σ : Type} → {m : Type → Type} → [self : MonadStateOf σ m] → σ → m PUnit
set
(σ :=
α: Type
α
) ∘ₑ
cast: {α β : Type} → α = β → α → β
cast
(
List.get_last: ∀ {α : Type 1} {a : α} {as : List α} {i : Fin (List.length (as ++ [a]))}, ¬i.val < List.length as → List.get (as ++ [a]) i = a
List.get_last
h: ¬x.val < List.length Δ
h
) ∘ₑ
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → List.get (List.append Δ [α]) x
e
) /-(S6)-/ |
Stmt.ite: {m : Type → Type ?u.190593} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg (List.append Δ [α]) → Bool
e
s₁: Stmt m ω Γ (List.append Δ [α]) b c β
s₁
s₂: Stmt m ω Γ (List.append Δ [α]) b c β
s₂
=>
Stmt.ite: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
(
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → Bool
e
) (
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s₁: Stmt m ω Γ (List.append Δ [α]) b c β
s₁
) (
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s₂: Stmt m ω Γ (List.append Δ [α]) b c β
s₂
) -- unreachable case; could be eliminated by a more precise specification of `ω`, but the benefit would be minimal |
Stmt.ret: {m : Type → Type ?u.190711} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg (List.append Δ [α]) → ω
e
=>
Stmt.ret: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
(
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → ω
e
) /-(S7)-/ |
Stmt.sbreak: {m : Type → Type ?u.190795} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
=>
Stmt.sbreak: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
/-(S8)-/ |
Stmt.scont: {m : Type → Type ?u.190845} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
=>
Stmt.scont: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
/-(S9)-/ |
Stmt.sfor: {m : Type → Type ?u.190894} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg (List.append Δ [α]) → List α✝
e
s: Stmt m ω (α✝ :: Γ) (List.append Δ [α]) true true Unit
s
=>
Stmt.sfor: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
(
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
e: Assg Γ → Assg (List.append Δ [α]) → List α✝
e
) (
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
get: {σ : Type} → {m : Type → Type} → [self : MonadState σ m] → m σ
get
)) (
Stmt.mapAssg: {Γ' Γ : List Type} → {m : Type → Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
Stmt.mapAssg
shadowSnd: {β : Type} → Assg (α :: β :: α :: Γ) → Assg (α :: β :: Γ)
shadowSnd
(
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s: Stmt m ω (α✝ :: Γ) (List.append Δ [α]) true true Unit
s
))) where @[simp]
unmut: {β : Type} → (Assg Γ → Assg (List.append Δ [α]) → β) → Assg (α :: Γ) → Assg Δ → β
unmut
{
β: Type
β
} (
e: Assg Γ → Assg (List.append Δ [α]) → β
e
:
Γ: List Type
Γ
Δ: List Type
Δ
++ [
α: Type
α
]
β: Type
β
) :
α: Type
α
::
Γ: List Type
Γ
Δ: List Type
Δ
β: Type
β
|
y: α
y
::
ρ: HList Γ
ρ
,
σ: Assg Δ
σ
=>
e: Assg Γ → Assg (List.append Δ [α]) → β
e
ρ: HList Γ
ρ
(
Assg.extendBot: {α : Type} → α → {Γ : List Type} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
y: α
y
σ: Assg Δ
σ
) @[simp]
shadowSnd: {Γ : List Type} → {α β : Type} → Assg (α :: β :: α :: Γ) → Assg (α :: β :: Γ)
shadowSnd
{
β: Type
β
} :
Assg: List Type → Type
Assg
(
α: Type
α
::
β: Type
β
::
α: Type
α
::
Γ: List Type
Γ
)
Assg: List Type → Type
Assg
(
α: Type
α
::
β: Type
β
::
Γ: List Type
Γ
) |
a': α
a'
::
b: β
b
:: _ ::
ρ: HList Γ
ρ
=>
a': α
a'
::
b: β
b
::
ρ: HList Γ
ρ
@[simp] def
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
[
Monad: (Type ?u.239464 → Type ?u.239463) → Type (max (?u.239464 + 1) ?u.239463)
Monad
m: Type → Type u_1
m
] :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
(
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
ω: Type
ω
m: Type → Type u_1
m
)
Empty: Type
Empty
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
/-(R1)-/ |
Stmt.ret: {m : Type → Type ?u.242023} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
∘ₑ
e: Assg Γ → Assg Δ → ω
e
) /-(R2)-/ |
Stmt.expr: {m : Type → Type ?u.242193} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg Δ → m α
e
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
∘ₑ
e: Assg Γ → Assg Δ → m α
e
) /-(R3)-/ |
Stmt.bind: {m : Type → Type ?u.242398} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s: Stmt m ω Γ Δ b c α✝
s
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
=>
Stmt.bind: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s: Stmt m ω Γ Δ b c α✝
s
) (
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
) /-(R4)-/ |
Stmt.letmut: {m : Type → Type ?u.242495} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
=>
Stmt.letmut: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
(
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
) /-(R5)-/ |
Stmt.assg: {m : Type → Type ?u.242597} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
=>
Stmt.assg: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
/-(R6)-/ |
Stmt.ite: {m : Type → Type ?u.242681} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
=>
Stmt.ite: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
(
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s₁: Stmt m ω Γ Δ b c α
s₁
) (
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s₂: Stmt m ω Γ Δ b c α
s₂
) /-(R7)-/ |
Stmt.sbreak: {m : Type → Type ?u.242784} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
=>
Stmt.sbreak: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
/-(R8)-/ |
Stmt.scont: {m : Type → Type ?u.242831} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
=>
Stmt.scont: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
/-(R9)-/ |
Stmt.sfor: {m : Type → Type ?u.242877} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
Stmt.sfor: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
(
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
) @[simp] def
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
[
Monad: (Type ?u.255179 → Type ?u.255178) → Type (max (?u.255179 + 1) ?u.255178)
Monad
m: Type → Type u_1
m
] :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
(
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
Unit: Type
Unit
m: Type → Type u_1
m
)
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
/-(L1)-/ |
Stmt.sbreak: {m : Type → Type ?u.256306} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
=>
Stmt.sbreak: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
/-(L2)-/ |
Stmt.scont: {m : Type → Type ?u.256353} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
=>
Stmt.scont: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
/-(L3)-/ |
Stmt.expr: {m : Type → Type ?u.256399} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg Δ → m α
e
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
∘ₑ
e: Assg Γ → Assg Δ → m α
e
) /-(L4)-/ |
Stmt.bind: {m : Type → Type ?u.256604} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s: Stmt m ω Γ Δ b c α✝
s
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
=>
Stmt.bind: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω Γ Δ b c α✝
s
) (
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
) /-(L5)-/ |
Stmt.letmut: {m : Type → Type ?u.256701} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
=>
Stmt.letmut: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
) /-(L6)-/ |
Stmt.assg: {m : Type → Type ?u.256803} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
=>
Stmt.assg: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
/-(L7)-/ |
Stmt.ite: {m : Type → Type ?u.256899} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
=>
Stmt.ite: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s₁: Stmt m ω Γ Δ b c α
s₁
) (
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s₂: Stmt m ω Γ Δ b c α
s₂
) |
Stmt.ret: {m : Type → Type ?u.257003} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
=>
Stmt.ret: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
/-(L8)-/ |
Stmt.sfor: {m : Type → Type ?u.257074} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
Stmt.sfor: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
) @[simp] def
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
[
Monad: (Type ?u.268935 → Type ?u.268934) → Type (max (?u.268935 + 1) ?u.268934)
Monad
m: Type → Type u_1
m
] :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
(
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
Unit: Type
Unit
m: Type → Type u_1
m
)
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
false: Bool
false
c: Bool
c
α: Type
α
/-(B1)-/ |
Stmt.sbreak: {m : Type → Type ?u.270257} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
Stmt.sbreak
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
(): Unit
()
) /-(B2)-/ |
Stmt.scont: {m : Type → Type ?u.270390} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
=>
Stmt.scont: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
/-(B3)-/ |
Stmt.expr: {m : Type → Type ?u.270436} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg Δ → m α
e
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
∘ₑ
e: Assg Γ → Assg Δ → m α
e
) /-(B4)-/ |
Stmt.bind: {m : Type → Type ?u.270641} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s: Stmt m ω Γ Δ b c α✝
s
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
=>
Stmt.bind: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s: Stmt m ω Γ Δ b c α✝
s
) (
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s': Stmt m ω (α✝ :: Γ) Δ b c α
s'
) /-(B5)-/ |
Stmt.letmut: {m : Type → Type ?u.270738} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
=>
Stmt.letmut: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
(
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
) /-(B6)-/ |
Stmt.assg: {m : Type → Type ?u.270840} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
=>
Stmt.assg: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
/-(B7)-/ |
Stmt.ite: {m : Type → Type ?u.270936} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
=>
Stmt.ite: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
(
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s₁: Stmt m ω Γ Δ b c α
s₁
) (
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s₂: Stmt m ω Γ Δ b c α
s₂
) |
Stmt.ret: {m : Type → Type ?u.271040} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
=>
Stmt.ret: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
/-(B8)-/ |
Stmt.sfor: {m : Type → Type ?u.271111} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
Stmt.sfor: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
) -- (elided in the paper) @[simp] def
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
[
Monad: (Type ?u.277926 → Type ?u.277925) → Type (max (?u.277926 + 1) ?u.277925)
Monad
m: Type → Type u_1
m
] :
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
false: Bool
false
c: Bool
c
α: Type
α
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
(
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
Unit: Type
Unit
m: Type → Type u_1
m
)
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
false: Bool
false
false: Bool
false
α: Type
α
|
Stmt.scont: {m : Type → Type ?u.278871} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
Stmt.scont
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
(): Unit
()
) |
Stmt.expr: {m : Type → Type ?u.278996} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg Δ → m α
e
=>
Stmt.expr: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
∘ₑ
e: Assg Γ → Assg Δ → m α
e
) |
Stmt.bind: {m : Type → Type ?u.279194} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s: Stmt m ω Γ Δ false c α✝
s
s': Stmt m ω (α✝ :: Γ) Δ false c α
s'
=>
Stmt.bind: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s: Stmt m ω Γ Δ false c α✝
s
) (
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s': Stmt m ω (α✝ :: Γ) Δ false c α
s'
) |
Stmt.letmut: {m : Type → Type ?u.279281} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
s: Stmt m ω Γ (α✝ :: Δ) false c α
s
=>
Stmt.letmut: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg Δ → α✝
e
(
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s: Stmt m ω Γ (α✝ :: Δ) false c α
s
) |
Stmt.assg: {m : Type → Type ?u.279374} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
=>
Stmt.assg: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
x: Fin (List.length Δ)
x
e: Assg Γ → Assg Δ → List.get Δ x
e
|
Stmt.ite: {m : Type → Type ?u.279463} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
s₁: Stmt m ω Γ Δ false c α
s₁
s₂: Stmt m ω Γ Δ false c α
s₂
=>
Stmt.ite: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg Δ → Bool
e
(
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s₁: Stmt m ω Γ Δ false c α
s₁
) (
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s₂: Stmt m ω Γ Δ false c α
s₂
) |
Stmt.ret: {m : Type → Type ?u.279557} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
=>
Stmt.ret: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg Δ → ω
e
|
Stmt.sfor: {m : Type → Type ?u.279621} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
Stmt.sfor: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg Δ → List α✝
e
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
)

The remaining function to be translated is D, which is straightforward as well except for its termination proof, as it recurses on the results of S (D3) and C ∘ B (D5). Because of rules (S2, S9) that introduce new bindings, S may in fact increase the size of the input, and the same is true for C and B for the sizeOf function automatically generated by Lean. Thus we introduce a new measure numExts that counts the number of special statements on top of basic do notation and prove that all three functions do not increase the size according to that measure. Because the rules (D3) and (D5) each eliminate such a special statement, it follows that D terminates because either the number of special statements decreases in each case, or it remains the same and the total number of statements decreases.

@[simp] def 
Stmt.numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
Stmt.numExts
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
Nat: Type
Nat
|
expr: {m : Type → Type ?u.314493} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
expr
_ =>
0: Nat
0
|
bind: {m : Type → Type ?u.314556} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
bind
s₁: Stmt m ω Γ Δ b c α✝
s₁
s₂: Stmt m ω (α✝ :: Γ) Δ b c α
s₂
=>
s₁: Stmt m ω Γ Δ b c α✝
s₁
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
+
s₂: Stmt m ω (α✝ :: Γ) Δ b c α
s₂
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
|
letmut: {m : Type → Type ?u.314678} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
letmut
_
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
=>
s: Stmt m ω Γ (α✝ :: Δ) b c α
s
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
+
1: Nat
1
|
assg: {m : Type → Type ?u.314809} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
assg
_ _ =>
1: Nat
1
|
ite: {m : Type → Type ?u.314889} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
ite
_
s₁: Stmt m ω Γ Δ b c α
s₁
s₂: Stmt m ω Γ Δ b c α
s₂
=>
s₁: Stmt m ω Γ Δ b c α
s₁
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
+
s₂: Stmt m ω Γ Δ b c α
s₂
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
|
ret: {m : Type → Type ?u.315027} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
ret
_ =>
1: Nat
1
|
sfor: {m : Type → Type ?u.315088} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
sfor
_
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
=>
s: Stmt m ω (α✝ :: Γ) Δ true true Unit
s
.
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
+
1: Nat
1
|
sbreak: {m : Type → Type ?u.315212} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → Stmt m ω Γ Δ true c α
sbreak
=>
1: Nat
1
|
scont: {m : Type → Type ?u.315254} → {ω : Type} → {Γ Δ : List Type} → {b : Bool} → {α : Type} → Stmt m ω Γ Δ b true α
scont
=>
1: Nat
1
@[simp] theorem
Stmt.numExts_mapAssg: ∀ {Γ' Γ : List Type} {m : Type → Type u_1} {ω : Type} {Δ : List Type} {b c : Bool} {β : Type} (f : Assg Γ' → Assg Γ) (s : Stmt m ω Γ Δ b c β), numExts (mapAssg f s) = numExts s
Stmt.numExts_mapAssg
(
f: Assg Γ' → Assg Γ
f
:
Assg: List Type → Type
Assg
Γ': List Type
Γ'
Assg: List Type → Type
Assg
Γ: List Type
Γ
) (
s: Stmt m ω Γ Δ b c β
s
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
) :
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
(
mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
mapAssg
f: Assg Γ' → Assg Γ
f
s: Stmt m ω Γ Δ b c β
s
) =
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
s: Stmt m ω Γ Δ b c β
s
:=
Γ', Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
f: Assg Γ' Assg Γ
s: Stmt m ω Γ Δ b c β

numExts (mapAssg f s) = numExts s
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α
Γ': List Type
f: Assg Γ' Assg Γ✝

expr
numExts (mapAssg f (expr e)) = numExts (expr e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (bind s s')) = numExts (bind s s')
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (letmut e s)) = numExts (letmut e s)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (assg x e)) = numExts (assg x e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (ite e s₁ s₂)) = numExts (ite e s₁ s₂)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (ret e)) = numExts (ret e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (sfor e s)) = numExts (sfor e s)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f sbreak) = numExts sbreak
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f scont) = numExts scont
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α
Γ': List Type
f: Assg Γ' Assg Γ✝

expr
numExts (mapAssg f (expr e)) = numExts (expr e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (bind s s')) = numExts (bind s s')
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (letmut e s)) = numExts (letmut e s)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (assg x e)) = numExts (assg x e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (ite e s₁ s₂)) = numExts (ite e s₁ s₂)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (ret e)) = numExts (ret e)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f (sfor e s)) = numExts (sfor e s)
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f sbreak) = numExts sbreak
Γ: List Type
m: Type Type u_1
ω: Type
Δ: List Type
b, c: Bool
β: Type
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
Γ': List Type
f: Assg Γ' Assg Γ✝
numExts (mapAssg f scont) = numExts scont

Goals accomplished! 🐙
theorem
Stmt.numExts_S: ∀ {m : Type → Type} {ω : Type} {Γ Δ : List Type} {α : Type} {b c : Bool} {β : Type} [inst : Monad m] (s : Stmt m ω Γ (List.append Δ [α]) b c β), numExts (S s) ≤ numExts s
Stmt.numExts_S
[
Monad: (Type ?u.330075 → Type ?u.330074) → Type (max (?u.330075 + 1) ?u.330074)
Monad
m: Type → Type
m
] (
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
:
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
m: Type → Type
m
ω: Type
ω
Γ: List Type
Γ
(
Δ: List Type
Δ
++ [
α: Type
α
])
b: Bool
b
c: Bool
c
β: Type
β
) :
numExts: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
(
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
)
numExts: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
:=
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
s: Stmt m ω Γ (List.append Δ [α]) b c β

numExts (S s) numExts s
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m

(s : Stmt m ω Γ (List.append Δ [α]) b c β), numExts (S s) numExts s
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m

{Δ' : List Type} (s : Stmt m ω Γ Δ' b c β) (h : Δ' = List.append Δ [α]), numExts (S (h s)) numExts s
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ': List Type
s: Stmt m ω Γ Δ' b c β
h: Δ' = List.append Δ [α]

numExts (S (h s)) numExts s
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ': List Type
s: Stmt m ω Γ Δ' b c β
h: Δ' = List.append Δ [α]

numExts (S (h s)) numExts s
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
Δ: List Type
e✝: Assg Γ✝ Assg (List.append Δ [α]) α
s✝: Stmt m ω Γ✝ :: List.append Δ [α]) b c β
ih: {Δ_1 : List Type} (h : α :: List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s)) numExts s

letmut
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) letmut e s)) numExts (letmut e s)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
Δ: List Type
s✝: Stmt m ω Γ✝ (List.append Δ [α]) b c α
s'✝: Stmt m ω (α :: Γ✝) (List.append Δ [α]) b c β
ih₁: {Δ_1 : List Type} (h : List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s)) numExts s
ih₂: {Δ_1 : List Type} (h : List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s')) numExts s'

bind
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) bind s s')) numExts (bind s s')

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
Δ: List Type
e✝: Assg Γ✝ Assg (List.append Δ [α]) α
s✝: Stmt m ω Γ✝ :: List.append Δ [α]) b c β
ih: {Δ_1 : List Type} (h : α :: List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s)) numExts s

letmut
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) letmut e s)) numExts (letmut e s)

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
Δ: List Type
x✝: Fin (List.length (List.append Δ [α]))
e✝: Assg Γ✝ Assg (List.append Δ [α]) List.get (List.append Δ [α]) x

assg
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) assg x e)) numExts (assg x e)

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝: Type
Δ: List Type
e✝: Assg Γ✝ Assg (List.append Δ [α]) Bool
s₁✝, s₂✝: Stmt m ω Γ✝ (List.append Δ [α]) b c α
ih₁: {Δ_1 : List Type} (h : List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s₁)) numExts s₁
ih₂: {Δ_1 : List Type} (h : List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s₂)) numExts s₂

ite
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) ite e s₁ s₂)) numExts (ite e s₁ s₂)

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
e✝: Assg Γ✝ Assg (List.append Δ [α]) List α
s✝: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: {Δ_1 : List Type} (h : List.append Δ [α] = List.append Δ_1 [α]), numExts (S (h s)) numExts s

sfor
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) sfor e s)) numExts (sfor e s)

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝: Type
Δ: List Type
e✝: Assg Γ✝ Assg (List.append Δ [α]) ω

ret
numExts (S ((_ : List.append Δ [α] = List.append Δ [α]) ret e)) numExts (ret e)

Goals accomplished! 🐙
theorem
Stmt.numExts_L_L: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {b c : Bool} {β : Type} [inst : Monad m] (s : Stmt m ω Γ Δ b c β), numExts (L (L s)) ≤ numExts s
Stmt.numExts_L_L
[
Monad: (Type ?u.339386 → Type ?u.339385) → Type (max (?u.339386 + 1) ?u.339385)
Monad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ b c β
s
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
) :
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
(
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω Γ Δ b c β
s
))
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
s: Stmt m ω Γ Δ b c β
s
:=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
s: Stmt m ω Γ Δ b c β

numExts (L (L s)) numExts s
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α

expr
numExts (L (L (expr e))) numExts (expr e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
numExts (L (L (bind s s'))) numExts (bind s s')
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
numExts (L (L (letmut e s))) numExts (letmut e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
numExts (L (L (assg x e))) numExts (assg x e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
numExts (L (L (ite e s₁ s₂))) numExts (ite e s₁ s₂)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
numExts (L (L (ret e))) numExts (ret e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
numExts (L (L (sfor e s))) numExts (sfor e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
numExts (L (L sbreak)) numExts sbreak
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
numExts (L (L scont)) numExts scont
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α

expr
numExts (L (L (expr e))) numExts (expr e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
numExts (L (L (bind s s'))) numExts (bind s s')
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
numExts (L (L (letmut e s))) numExts (letmut e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
numExts (L (L (assg x e))) numExts (assg x e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
numExts (L (L (ite e s₁ s₂))) numExts (ite e s₁ s₂)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
numExts (L (L (ret e))) numExts (ret e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
numExts (L (L (sfor e s))) numExts (sfor e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
numExts (L (L sbreak)) numExts sbreak
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
numExts (L (L scont)) numExts scont

Goals accomplished! 🐙
theorem
Stmt.numExts_C_B: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {b c : Bool} {β : Type} [inst : Monad m] (s : Stmt m ω Γ Δ b c β), numExts (C (B s)) ≤ numExts s
Stmt.numExts_C_B
[
Monad: (Type ?u.342450 → Type ?u.342449) → Type (max (?u.342450 + 1) ?u.342449)
Monad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ b c β
s
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
) :
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
(
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
(
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s: Stmt m ω Γ Δ b c β
s
))
numExts: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
s: Stmt m ω Γ Δ b c β
s
:=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
s: Stmt m ω Γ Δ b c β

numExts (C (B s)) numExts s
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α

expr
numExts (C (B (expr e))) numExts (expr e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
numExts (C (B (bind s s'))) numExts (bind s s')
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
numExts (C (B (letmut e s))) numExts (letmut e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
numExts (C (B (assg x e))) numExts (assg x e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
numExts (C (B (ite e s₁ s₂))) numExts (ite e s₁ s₂)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
numExts (C (B (ret e))) numExts (ret e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
numExts (C (B (sfor e s))) numExts (sfor e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
numExts (C (B sbreak)) numExts sbreak
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
numExts (C (B scont)) numExts scont
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ m α

expr
numExts (C (B (expr e))) numExts (expr e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
numExts (C (B (bind s s'))) numExts (bind s s')
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
numExts (C (B (letmut e s))) numExts (letmut e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
numExts (C (B (assg x e))) numExts (assg x e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ Bool
s₁✝, s₂✝: Stmt m ω Γ✝ Δ✝ b c α
numExts (C (B (ite e s₁ s₂))) numExts (ite e s₁ s₂)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝: Type
e✝: Assg Γ✝ Assg Δ✝ ω
numExts (C (B (ret e))) numExts (ret e)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e✝: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
numExts (C (B (sfor e s))) numExts (sfor e s)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
numExts (C (B sbreak)) numExts sbreak
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
β: Type
inst✝: Monad m
Γ✝, Δ✝: List Type
b✝: Bool
α✝: Type
numExts (C (B scont)) numExts scont

Goals accomplished! 🐙
-- Auxiliary tactic for showing that `D` terminates macro "D_tac" : tactic => `({simp_wf solve | apply Prod.Lex.left; assumption | apply Prod.Lex.right' <;> simp_arith }) @[simp] def
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
[
Monad: (Type ?u.348150 → Type ?u.348149) → Type (max (?u.348150 + 1) ?u.348149)
Monad
m: Type → Type
m
] :
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
m: Type → Type
m
Empty: Type
Empty
Γ: List Type
Γ
: List Type
false: Bool
false
false: Bool
false
α: Type
α
(
Γ: List Type
Γ
m: Type → Type
m
α: Type
α
) |
Stmt.expr: {m : Type → Type ?u.349536} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg ∅ → m α
e
=> (
e: Assg Γ → Assg ∅ → m α
e
[·][
: Assg ∅
]) |
Stmt.bind: {m : Type → Type ?u.349597} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s: Stmt m Empty Γ ∅ false false α✝
s
s': Stmt m Empty (α✝ :: Γ) ∅ false false α
s'
=> (fun
ρ: Assg Γ
ρ
=>
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
s: Stmt m Empty Γ ∅ false false α✝
s
ρ: Assg Γ
ρ
>>= fun
x: α✝
x
=>
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
s': Stmt m Empty (α✝ :: Γ) ∅ false false α
s'
(
x: α✝
x
::
ρ: Assg Γ
ρ
)) |
Stmt.letmut: {m : Type → Type ?u.349872} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg ∅ → α✝
e
s: Stmt m Empty Γ (α✝ :: ∅) false false α
s
=> have :=
Nat.lt_succ_of_le: ∀ {n m : Nat}, n ≤ m → n < Nat.succ m
Nat.lt_succ_of_le
<|
Stmt.numExts_S: ∀ {m : Type → Type} {ω : Type} {Γ Δ : List Type} {α : Type} {b c : Bool} {β : Type} [inst : Monad m] (s : Stmt m ω Γ (List.append Δ [α]) b c β), Stmt.numExts (S s) ≤ Stmt.numExts s
Stmt.numExts_S
(Δ :=
[]: List Type
[]
)
s: Stmt m Empty Γ (α✝ :: ∅) false false α
s
-- for termination fun
ρ: Assg Γ
ρ
=> let
x: α✝
x
:=
e: Assg Γ → Assg ∅ → α✝
e
[
ρ: Assg Γ
ρ
][
: Assg ∅
]
StateT.run': {σ : Type} → {m : Type → Type} → [inst : Functor m] → {α : Type} → StateT σ m α → σ → m α
StateT.run'
(
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
(
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s: Stmt m Empty Γ (α✝ :: ∅) false false α
s
) (
x: α✝
x
::
ρ: Assg Γ
ρ
))
x: α✝
x
|
Stmt.ite: {m : Type → Type ?u.350351} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg ∅ → Bool
e
s₁: Stmt m Empty Γ ∅ false false α
s₁
s₂: Stmt m Empty Γ ∅ false false α
s₂
=> (fun
ρ: Assg Γ
ρ
=> if
e: Assg Γ → Assg ∅ → Bool
e
[
ρ: Assg Γ
ρ
][
: Assg ∅
] then
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
s₁: Stmt m Empty Γ ∅ false false α
s₁
ρ: Assg Γ
ρ
else
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
s₂: Stmt m Empty Γ ∅ false false α
s₂
ρ: Assg Γ
ρ
) |
Stmt.sfor: {m : Type → Type ?u.350515} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg ∅ → List α✝
e
s: Stmt m Empty (α✝ :: Γ) ∅ true true Unit
s
=> have :=
Nat.lt_succ_of_le: ∀ {n m : Nat}, n ≤ m → n < Nat.succ m
Nat.lt_succ_of_le
<|
Stmt.numExts_C_B: ∀ {m : Type → Type} {ω : Type} {Γ Δ : List Type} {b c : Bool} {β : Type} [inst : Monad m] (s : Stmt m ω Γ Δ b c β), Stmt.numExts (C (B s)) ≤ Stmt.numExts s
Stmt.numExts_C_B
(Δ :=
[]: List Type
[]
)
s: Stmt m Empty (α✝ :: Γ) ∅ true true Unit
s
-- for termination fun
ρ: Assg Γ
ρ
=>
runCatch: {m : Type → Type} → {α : Type} → [inst : Monad m] → ExceptT α m α → m α
runCatch
(
forM: {m : Type → Type} → {γ α : Type} → [self : ForM m γ α] → [inst : Monad m] → γ → (α → m PUnit) → m PUnit
forM
e: Assg Γ → Assg ∅ → List α✝
e
[
ρ: Assg Γ
ρ
][
: Assg ∅
] (fun
x: α✝
x
=>
runCatch: {m : Type → Type} → {α : Type} → [inst : Monad m] → ExceptT α m α → m α
runCatch
(
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
(
C: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
(
B: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s: Stmt m Empty (α✝ :: Γ) ∅ true true Unit
s
)) (
x: α✝
x
::
ρ: Assg Γ
ρ
)))) |
Stmt.ret: {m : Type → Type ?u.351104} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg ∅ → Empty
e
=> (nomatch
e: Assg Γ → Assg ∅ → Empty
e
[·][
: Assg ∅
]) termination_by _ s => (
s: Stmt m Empty Γ ∅ false false α
s
.
numExts: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
,
sizeOf: {α : Type 1} → [self : SizeOf α] → α → Nat
sizeOf
s: Stmt m Empty Γ ∅ false false α
s
) decreasing_by

Goals accomplished! 🐙

Finally we compose D and R into the translation rule for a top-level statement (1').

def 
Do.trans: {m : Type → Type} → {α : Type} → [inst : Monad m] → Do m α → m α
Do.trans
[
Monad: (Type → Type) → Type 1
Monad
m: Type → Type
m
] (
s: Do m α
s
:
Do: (Type → Type) → Type → Type 1
Do
m: Type → Type
m
α: Type
α
) :
m: Type → Type
m
α: Type
α
:=
runCatch: {m : Type → Type} → {α : Type} → [inst : Monad m] → ExceptT α m α → m α
runCatch
(
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
(
R: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s: Do m α
s
)
: Assg ∅
)

Equivalence Proof

Using the monadic dynamic semantics, we can modularly prove for each individual translation function that evaluating its output is equivalent to directly evaluating the input, modulo some lifting and adjustment of resulting values. After induction on the statement, the proofs are mostly concerned with case splitting, application of congruence theorems, and simplification. We can mostly offload these tasks onto Aesop.

attribute [local simp] 
map_eq_pure_bind: ∀ {m : Type u_1 → Type u_2} {α β : Type u_1} [inst : Monad m] [inst_1 : LawfulMonad m] (f : α → β) (x : m α), f <$> x = do let a ← x pure (f a)
map_eq_pure_bind
ExceptT.run_bind: ∀ {m : Type u_1 → Type u_2} {ε α β : Type u_1} {f : α → ExceptT ε m β} [inst : Monad m] (x : ExceptT ε m α), ExceptT.run (x >>= f) = do let x ← ExceptT.run x match x with | Except.ok x => ExceptT.run (f x) | Except.error e => pure (Except.error e)
ExceptT.run_bind
attribute [aesop safe apply]
bind_congr: ∀ {m : Type u_1 → Type u_2} {α β : Type u_1} [inst : Bind m] {x : m α} {f g : α → m β}, (∀ (a : α), f a = g a) → x >>= f = x >>= g
bind_congr
theorem
eval_R: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {b c : Bool} {α : Type} {ρ : Assg Γ} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Stmt m ω Γ Δ b c α), Stmt.eval ρ (R s) σ = do let x ← ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
eval_R
[
Monad: (Type ?u.373143 → Type ?u.373142) → Type (max (?u.373143 + 1) ?u.373142)
Monad
m: Type → Type u_1
m
] [
LawfulMonad: (m : Type ?u.373172 → Type ?u.373171) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ b c α
s
:
Stmt: (Type → Type ?u.373206) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 ?u.373206)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
) : (
R: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT ω m) Empty Γ Δ b c α
R
s: Stmt m ω Γ Δ b c α
s
).
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
= (
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
(
s: Stmt m ω Γ Δ b c α
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
) >>= fun
x: Neut ω α b c × Assg Δ
x
=> match (generalizing := false)
x: Neut ω α b c × Assg Δ
x
with | (
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
, _) =>
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
o: ω
o
| (
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) | (
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ: Assg Δ
σ
) | (
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
,
σ: Assg Δ
σ
) :
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
ω: Type
ω
m: Type → Type u_1
m
(
Neut: Type → Type → Bool → Bool → Type
Neut
Empty: Type
Empty
α: Type
α
b: Bool
b
c: Bool
c
×
Assg: List Type → Type
Assg
Δ: List Type
Δ
)) :=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

Stmt.eval ρ (R s) σ = do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (R s) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (R s) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval ρ (R (Stmt.sfor e s)) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ (Stmt.sfor e s) σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval.go ρ c b α (R s) σ (e ρ σ)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (e ρ σ)) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c b α (R s) σ []) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ []) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c b α (R s) σ (head :: tail)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c b α (R s) σ []) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ []) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c b α (R s) σ (head :: tail)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)

Goals accomplished! 🐙
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
ρ: Assg Γ✝
σ: Assg Δ✝

h.letmut
ExceptT.run (Stmt.eval ρ (R (Stmt.letmut e s)) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ (Stmt.letmut e s) σ) match b, c, x with | b, c, (Neut.ret o, snd) => throw o | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, σ) => pure (Neut.rbreak, σ)

Goals accomplished! 🐙
@[simp] theorem
eval_mapAssg: ∀ {m : Type → Type u_1} {Γ' Γ : List Type} {ω : Type} {Δ : List Type} {b c : Bool} {β : Type} {ρ : Assg Γ'} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (f : Assg Γ' → Assg Γ) (s : Stmt m ω Γ Δ b c β), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
eval_mapAssg
[
Monad: (Type ?u.513583 → Type ?u.513582) → Type (max (?u.513583 + 1) ?u.513582)
Monad
m: Type → Type u_1
m
] [
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
f: Assg Γ' → Assg Γ
f
:
Assg: List (Type ?u.513856) → Type ?u.513856
Assg
Γ': List Type
Γ'
Assg: List (Type ?u.513712) → Type ?u.513712
Assg
Γ: List Type
Γ
) : (
s: Stmt m ω Γ Δ b c β
s
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
β: Type
β
),
Stmt.eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
Stmt.eval
ρ: Assg Γ'
ρ
(
Stmt.mapAssg: {Γ' Γ : List Type} → {m : Type → Type u_1} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {β : Type} → (Assg Γ' → Assg Γ) → Stmt m ω Γ Δ b c β → Stmt m ω Γ' Δ b c β
Stmt.mapAssg
f: Assg Γ' → Assg Γ
f
s: Stmt m ω Γ Δ b c β
s
)
σ: Assg Δ
σ
=
Stmt.eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
Stmt.eval
(
f: Assg Γ' → Assg Γ
f
ρ: Assg Γ'
ρ
)
s: Stmt m ω Γ Δ b c β
s
σ: Assg Δ
σ
:=
m: Type Type u_1
Γ', Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
ρ: Assg Γ'
σ: Assg Δ
inst✝¹: Monad m
f: Assg Γ' Assg Γ

(s : Stmt m ω Γ Δ b c β), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
m: Type Type u_1
Γ', Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
ρ: Assg Γ'
σ: Assg Δ
inst✝¹: Monad m
f: Assg Γ' Assg Γ
s: Stmt m ω Γ Δ b c β

Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
m: Type Type u_1
Γ', Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
ρ: Assg Γ'
σ: Assg Δ
inst✝¹: Monad m
f: Assg Γ' Assg Γ
s: Stmt m ω Γ Δ b c β

Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
σ: Assg Δ✝
f: Assg Γ' Assg Γ✝

sfor
Stmt.eval ρ (Stmt.mapAssg f (Stmt.sfor e s)) σ = Stmt.eval (f ρ) (Stmt.sfor e s) σ
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
σ: Assg Δ✝
f: Assg Γ' Assg Γ✝

sfor
Stmt.eval.go ρ c b α (Stmt.mapAssg (fun x => match x with | (a, as) => a :: f as) s) σ (e (f ρ) σ) = Stmt.eval.go (f ρ) c b α s σ (e (f ρ) σ)
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
f: Assg Γ' Assg Γ✝
σ: Assg Δ✝

sfor.nil
Stmt.eval.go ρ c b α (Stmt.mapAssg (fun x => match x with | (a, as) => a :: f as) s) σ [] = Stmt.eval.go (f ρ) c b α s σ []
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
f: Assg Γ' Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
Stmt.eval.go ρ c b α (Stmt.mapAssg (fun x => match x with | (a, as) => a :: f as) s) σ (head :: tail) = Stmt.eval.go (f ρ) c b α s σ (head :: tail)
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
f: Assg Γ' Assg Γ✝
σ: Assg Δ✝

sfor.nil
Stmt.eval.go ρ c b α (Stmt.mapAssg (fun x => match x with | (a, as) => a :: f as) s) σ [] = Stmt.eval.go (f ρ) c b α s σ []
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ih: {Γ' : List Type} {ρ : Assg Γ'} {σ : Assg Δ✝} (f : Assg Γ' Assg (α :: Γ✝)), Stmt.eval ρ (Stmt.mapAssg f s) σ = Stmt.eval (f ρ) s σ
Γ': List Type
ρ: Assg Γ'
f: Assg Γ' Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
Stmt.eval.go ρ c b α (Stmt.mapAssg (fun x => match x with | (a, as) => a :: f as) s) σ (head :: tail) = Stmt.eval.go (f ρ) c b α s σ (head :: tail)

Goals accomplished! 🐙
m: Type Type u_1
Γ: List Type
ω: Type
Δ: List Type
b, c: Bool
β: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
Γ': List Type
ρ: Assg Γ'
σ: Assg Δ✝
f: Assg Γ' Assg Γ✝

letmut
Stmt.eval ρ (Stmt.mapAssg f (Stmt.letmut e s)) σ = Stmt.eval (f ρ) (Stmt.letmut e s) σ

Goals accomplished! 🐙

We need one last helper function on context bottoms to be able to state the invariant of S, and then prove various lemmas about their interactions.

def 
Assg.bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
Assg.bot
: {
Γ: List (Type u_1)
Γ
:
_: Type (u_1 + 1)
_
}
Assg: List (Type u_1) → Type u_1
Assg
(
Γ: List (Type u_1)
Γ
++ [
α: Type u_1
α
])
α: Type u_1
α
| [], [
a: α
a
] =>
a: α
a
| _ ::
_: List (Type u_1)
_
, _ ::
as: HList (List.append tail✝ [α])
as
=>
bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
bot
as: HList (List.append tail✝ [α])
as
@[simp] theorem
Assg.dropBot_extendBot: ∀ {α : Type u_1} {Γ : List (Type u_1)} (a : α) (σ : Assg Γ), dropBot (extendBot a σ) = σ
Assg.dropBot_extendBot
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
) :
Assg.dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
(
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
) =
σ: Assg Γ
σ
:=
α: Type u_1
Γ: List (Type u_1)
a: α
σ: Assg Γ

dropBot (extendBot a σ) = σ
α: Type u_1
a: α
σ: Assg []

nil
dropBot (extendBot a σ) = σ
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (head :: tail)
dropBot (extendBot a σ) = σ
α: Type u_1
a: α
σ: Assg []

nil
dropBot (extendBot a σ) = σ
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (head :: tail)
dropBot (extendBot a σ) = σ
α: Type u_1
a: α

nil.unit
dropBot (extendBot a PUnit.unit) = PUnit.unit
α: Type u_1
a: α

nil.unit
dropBot (extendBot a PUnit.unit) = PUnit.unit

Goals accomplished! 🐙
@[simp] theorem
Assg.bot_extendBot: ∀ {α : Type u_1} {Γ : List (Type u_1)} (a : α) (σ : Assg Γ), bot (extendBot a σ) = a
Assg.bot_extendBot
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
) :
Assg.bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
Assg.bot
(
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
) =
a: α
a
:=
α: Type u_1
Γ: List (Type u_1)
a: α
σ: Assg Γ

bot (extendBot a σ) = a
α: Type u_1
a: α
σ: Assg []

nil
bot (extendBot a σ) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (head :: tail)
bot (extendBot a σ) = a
α: Type u_1
a: α
σ: Assg []

nil
bot (extendBot a σ) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (head :: tail)
bot (extendBot a σ) = a
α: Type u_1
a: α

nil.unit
bot (extendBot a PUnit.unit) = a
α: Type u_1
a: α

nil.unit
bot (extendBot a PUnit.unit) = a

Goals accomplished! 🐙
@[simp] theorem
Assg.extendBot_bot_dropBot: ∀ {Γ : List (Type u_1)} {α : Type u_1} (σ : Assg (List.append Γ [α])), extendBot (bot σ) (dropBot σ) = σ
Assg.extendBot_bot_dropBot
(
σ: Assg (List.append Γ [α])
σ
:
Assg: List (Type u_1) → Type u_1
Assg
(
Γ: List (Type u_1)
Γ
++ [
α: Type u_1
α
])) :
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
(
Assg.bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
Assg.bot
σ: Assg (List.append Γ [α])
σ
) (
Assg.dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
σ: Assg (List.append Γ [α])
σ
) =
σ: Assg (List.append Γ [α])
σ
:=
Γ: List (Type u_1)
α: Type u_1
σ: Assg (List.append Γ [α])

extendBot (bot σ) (dropBot σ) = σ
α: Type u_1
σ: Assg (List.append [] [α])

nil
extendBot (bot σ) (dropBot σ) = σ
α, head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (List.append (head :: tail) [α])
extendBot (bot σ) (dropBot σ) = σ
α: Type u_1
σ: Assg (List.append [] [α])

nil
extendBot (bot σ) (dropBot σ) = σ
α, head✝: Type u_1
tail✝: List (Type u_1)
σ: Assg (List.append (head :: tail) [α])
extendBot (bot σ) (dropBot σ) = σ
α: Type u_1
fst✝: α
snd✝: HList []

nil.mk
extendBot (bot (fst, snd)) (dropBot (fst, snd)) = (fst, snd)
α: Type u_1
fst✝: α
snd✝: HList []

nil.mk
extendBot (bot (fst, snd)) (dropBot (fst, snd)) = (fst, snd)

Goals accomplished! 🐙

Goals accomplished! 🐙

Goals accomplished! 🐙
@[simp] theorem
Assg.dropBot_set_extendBot_init: ∀ {α : Type u_1} {Γ : List (Type u_1)} {i : Fin (List.length (List.append Γ [α]))} (a : α) (σ : Assg Γ) (h : i.val < List.length Γ) {b : List.get (List.append Γ [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (Γ ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (Γ ++ [α])) } = List.get Γ { val := i.val, isLt := h }) ▸ b)
Assg.dropBot_set_extendBot_init
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
) (
h: i.val < List.length Γ
h
:
i: Fin (List.length (List.append Γ [α]))
i
.
1: {n : Nat} → Fin n → Nat
1
<
Γ: List (Type u_1)
Γ
.
length: {α : Type (u_1 + 1)} → List α → Nat
length
) {
b: List.get (List.append Γ [α]) i
b
} :
Assg.dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
((
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
).
set: {αs : List (Type u_1)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
i: Fin (List.length (List.append Γ [α]))
i
b: List.get (List.append Γ [α]) i
b
) =
σ: Assg Γ
σ
.
set: {αs : List (Type u_1)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
i: Fin (List.length (List.append Γ [α]))
i
.
1: {n : Nat} → Fin n → Nat
1
,
h: i.val < List.length Γ
h
⟩ (
List.get_append_left: ∀ {α : Type (u_1 + 1)} {i : Nat} (as bs : List α) (h : i < List.length as) {h' : i < List.length (as ++ bs)}, List.get (as ++ bs) { val := i, isLt := h' } = List.get as { val := i, isLt := h }
List.get_append_left
..
b: List.get (List.append Γ [α]) i
b
) :=
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: i.val < List.length Γ
b: List.get (List.append Γ [α]) i

dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (Γ ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (Γ ++ [α])) } = List.get Γ { val := i.val, isLt := h }) b)
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: i.val < List.length Γ
b: List.get (List.append Γ [α]) i

dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (Γ ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (Γ ++ [α])) } = List.get Γ { val := i.val, isLt := h }) b)
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h: i.val < List.length []
b: List.get (List.append [] [α]) i

nil
dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get ([] ++ [α]) { val := i.val, isLt := (_ : i.val < List.length ([] ++ [α])) } = List.get [] { val := i.val, isLt := h }) b)

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
σ: Assg (head :: tail)
h: i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i

cons
dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := i.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
h: i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i
fst✝: head
snd✝: HList tail

cons.mk
dropBot (HList.set (extendBot a (fst, snd)) i b) = HList.set (fst, snd) { val := i.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := i.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i✝: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
i: Nat
h': i < List.length (List.append (head :: tail) [α])
h: { val := i, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := i, isLt := h' }

cons.mk
dropBot (HList.set (extendBot a (fst, snd)) { val := i, isLt := h' } b) = HList.set (fst, snd) { val := { val := i, isLt := h' }.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := { val := i, isLt := h' }.val, isLt := (_ : { val := i, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := i, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: { val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = HList.set (fst, snd) { val := { val := Nat.zero, isLt := h' }.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.zero, isLt := h' }.val, isLt := (_ : { val := Nat.zero, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.zero, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = HList.set (fst, snd) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.succ n, isLt := h' }.val, isLt := (_ : { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: { val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = HList.set (fst, snd) { val := { val := Nat.zero, isLt := h' }.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.zero, isLt := h' }.val, isLt := (_ : { val := Nat.zero, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.zero, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = HList.set (fst, snd) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.succ n, isLt := h' }.val, isLt := (_ : { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: dropBot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = fst :: HList.set snd { val := n, isLt := (_ : Nat.succ n List.length tail) } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.succ n, isLt := h' }.val, isLt := (_ : { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: dropBot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = fst :: HList.set snd { val := n, isLt := (_ : Nat.succ n List.length tail) } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.succ n, isLt := h' }.val, isLt := (_ : { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: HList.set snd { val := { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val, isLt := ?cons.mk.succ.h } ((_ : List.get (tail ++ [α]) { val := { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val, isLt := (_ : { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length (tail ++ [α])) } = List.get tail { val := { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val, isLt := ?cons.mk.succ.h }) b) = fst :: HList.set snd { val := n, isLt := (_ : Nat.succ n List.length tail) } ((_ : List.get (head :: tail ++ [α]) { val := { val := Nat.succ n, isLt := h' }.val, isLt := (_ : { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail ++ [α])) } = List.get (head :: tail) { val := { val := Nat.succ n, isLt := h' }.val, isLt := h }) b)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail) (h : i.val < List.length tail) {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = HList.set σ { val := i.val, isLt := h } ((_ : List.get (tail ++ [α]) { val := i.val, isLt := (_ : i.val < List.length (tail ++ [α])) } = List.get tail { val := i.val, isLt := h }) b)
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail

Goals accomplished! 🐙
@[simp] theorem
Assg.bot_set_extendBot_init: ∀ {α : Type u_1} {Γ : List (Type u_1)} {i : Fin (List.length (List.append Γ [α]))} (a : α) (σ : Assg Γ), i.val < List.length Γ → ∀ {b : List.get (List.append Γ [α]) i}, bot (HList.set (extendBot a σ) i b) = a
Assg.bot_set_extendBot_init
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
) (
h: i.val < List.length Γ
h
:
i: Fin (List.length (List.append Γ [α]))
i
.
1: {n : Nat} → Fin n → Nat
1
<
Γ: List (Type u_1)
Γ
.
length: {α : Type (u_1 + 1)} → List α → Nat
length
) {
b: List.get (List.append Γ [α]) i
b
} :
Assg.bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
Assg.bot
((
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
).
set: {αs : List (Type u_1)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
i: Fin (List.length (List.append Γ [α]))
i
b: List.get (List.append Γ [α]) i
b
) =
a: α
a
:=
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: i.val < List.length Γ
b: List.get (List.append Γ [α]) i

bot (HList.set (extendBot a σ) i b) = a
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: i.val < List.length Γ
b: List.get (List.append Γ [α]) i

bot (HList.set (extendBot a σ) i b) = a
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h: i.val < List.length []
b: List.get (List.append [] [α]) i

nil
bot (HList.set (extendBot a σ) i b) = a

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
σ: Assg (head :: tail)
h: i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i

cons
bot (HList.set (extendBot a σ) i b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
h: i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i
fst✝: head
snd✝: HList tail

cons.mk
bot (HList.set (extendBot a (fst, snd)) i b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i✝: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
i: Nat
h': i < List.length (List.append (head :: tail) [α])
h: { val := i, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := i, isLt := h' }

cons.mk
bot (HList.set (extendBot a (fst, snd)) { val := i, isLt := h' } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: { val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
bot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
bot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: { val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
bot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
bot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
bot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
bot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
a = a
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = a
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: { val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ.h
{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
;

Goals accomplished! 🐙
@[simp] theorem
Assg.dropBot_set_extendBot_bottom: ∀ {α : Type u_1} {Γ : List (Type u_1)} {i : Fin (List.length (List.append Γ [α]))} (a : α) (σ : Assg Γ), ¬i.val < List.length Γ → ∀ {b : List.get (List.append Γ [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
Assg.dropBot_set_extendBot_bottom
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type ?u.548037) → Type ?u.548037
Assg
Γ: List (Type u_1)
Γ
) (
h: ¬i.val < List.length Γ
h
: ¬
i: Fin (List.length (List.append Γ [α]))
i
.
1: {n : Nat} → Fin n → Nat
1
<
Γ: List (Type u_1)
Γ
.
length: {α : Type (u_1 + 1)} → List α → Nat
length
) {
b: List.get (List.append Γ [α]) i
b
} :
Assg.dropBot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
((
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
).
set: {αs : List (Type u_1)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
i: Fin (List.length (List.append Γ [α]))
i
b: List.get (List.append Γ [α]) i
b
) =
σ: Assg Γ
σ
:=
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: ¬i.val < List.length Γ
b: List.get (List.append Γ [α]) i

dropBot (HList.set (extendBot a σ) i b) = σ
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: ¬i.val < List.length Γ
b: List.get (List.append Γ [α]) i

dropBot (HList.set (extendBot a σ) i b) = σ
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h: ¬i.val < List.length []
b: List.get (List.append [] [α]) i

nil
dropBot (HList.set (extendBot a σ) i b) = σ

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
σ: Assg (head :: tail)
h: ¬i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i

cons
dropBot (HList.set (extendBot a σ) i b) = σ
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
h: ¬i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i
fst✝: head
snd✝: HList tail

cons.mk
dropBot (HList.set (extendBot a (fst, snd)) i b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i✝: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
i: Nat
h': i < List.length (List.append (head :: tail) [α])
h: ¬{ val := i, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := i, isLt := h' }

cons.mk
dropBot (HList.set (extendBot a (fst, snd)) { val := i, isLt := h' } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = (fst, snd)

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
dropBot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: dropBot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: dropBot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
fst :: snd = (fst, snd)
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
¬{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ.h
¬{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, dropBot (HList.set (extendBot a σ) i b) = σ
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
h'': { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail

cons.mk.succ.h
False

Goals accomplished! 🐙
@[simp] theorem
Assg.bot_set_extendBot_bottom: ∀ {α : Type u_1} {Γ : List (Type u_1)} {i : Fin (List.length (List.append Γ [α]))} (a : α) (σ : Assg Γ), ¬i.val < List.length Γ → ∀ {b : List.get (List.append Γ [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (Γ ++ [α]) i = α) b
Assg.bot_set_extendBot_bottom
(
a: α
a
:
α: Type u_1
α
) (
σ: Assg Γ
σ
:
Assg: List (Type u_1) → Type u_1
Assg
Γ: List (Type u_1)
Γ
) (
h: ¬i.val < List.length Γ
h
: ¬
i: Fin (List.length (List.append Γ [α]))
i
.
1: {n : Nat} → Fin n → Nat
1
<
Γ: List (Type u_1)
Γ
.
length: {α : Type (u_1 + 1)} → List α → Nat
length
) {
b: List.get (List.append Γ [α]) i
b
} :
Assg.bot: {α : Type u_1} → {Γ : List (Type u_1)} → Assg (List.append Γ [α]) → α
Assg.bot
((
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Γ
σ
).
set: {αs : List (Type u_1)} → HList αs → (i : Fin (List.length αs)) → List.get αs i → HList αs
set
i: Fin (List.length (List.append Γ [α]))
i
b: List.get (List.append Γ [α]) i
b
) =
cast: {α β : Type u_1} → α = β → α → β
cast
(
List.get_last: ∀ {α : Type (u_1 + 1)} {a : α} {as : List α} {i : Fin (List.length (as ++ [a]))}, ¬i.val < List.length as → List.get (as ++ [a]) i = a
List.get_last
h: ¬i.val < List.length Γ
h
)
b: List.get (List.append Γ [α]) i
b
:=
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: ¬i.val < List.length Γ
b: List.get (List.append Γ [α]) i

bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (Γ ++ [α]) i = α) b
α: Type u_1
Γ: List (Type u_1)
i: Fin (List.length (List.append Γ [α]))
a: α
σ: Assg Γ
h: ¬i.val < List.length Γ
b: List.get (List.append Γ [α]) i

bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (Γ ++ [α]) i = α) b
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h: ¬i.val < List.length []
b: List.get (List.append [] [α]) i

nil
bot (HList.set (extendBot a σ) i b) = cast (_ : List.get ([] ++ [α]) i = α) b
α: Type u_1
a: α
i✝: Fin (List.length (List.append [] [α]))
σ: Assg []
i: Nat
h': i < List.length (List.append [] [α])
h: ¬{ val := i, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := i, isLt := h' }

nil
bot (HList.set (extendBot a σ) { val := i, isLt := h' } b) = cast (_ : List.get ([] ++ [α]) { val := i, isLt := h' } = α) b
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h': Nat.zero < List.length (List.append [] [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := Nat.zero, isLt := h' }

nil.zero
bot (HList.set (extendBot a σ) { val := Nat.zero, isLt := h' } b) = cast (_ : List.get ([] ++ [α]) { val := Nat.zero, isLt := h' } = α) b
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
n✝: Nat
h': Nat.succ n < List.length (List.append [] [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := Nat.succ n, isLt := h' }
bot (HList.set (extendBot a σ) { val := Nat.succ n, isLt := h' } b) = cast (_ : List.get ([] ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h': Nat.zero < List.length (List.append [] [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := Nat.zero, isLt := h' }

nil.zero
bot (HList.set (extendBot a σ) { val := Nat.zero, isLt := h' } b) = cast (_ : List.get ([] ++ [α]) { val := Nat.zero, isLt := h' } = α) b
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
h': Nat.zero < List.length (List.append [] [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := Nat.zero, isLt := h' }

nil.zero
b = cast (_ : List.get ([] ++ [α]) { val := Nat.zero, isLt := h' } = α) b
;

Goals accomplished! 🐙
α: Type u_1
a: α
i: Fin (List.length (List.append [] [α]))
σ: Assg []
n✝: Nat
h': Nat.succ n < List.length (List.append [] [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length []
b: List.get (List.append [] [α]) { val := Nat.succ n, isLt := h' }

nil.succ
bot (HList.set (extendBot a σ) { val := Nat.succ n, isLt := h' } b) = cast (_ : List.get ([] ++ [α]) { val := Nat.succ n, isLt := h' } = α) b

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
σ: Assg (head :: tail)
h: ¬i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i

cons
bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (head :: tail ++ [α]) i = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
h: ¬i.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) i
fst✝: head
snd✝: HList tail

cons.mk
bot (HList.set (extendBot a (fst, snd)) i b) = cast (_ : List.get (head :: tail ++ [α]) i = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i✝: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
i: Nat
h': i < List.length (List.append (head :: tail) [α])
h: ¬{ val := i, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := i, isLt := h' }

cons.mk
bot (HList.set (extendBot a (fst, snd)) { val := i, isLt := h' } b) = cast (_ : List.get (head :: tail ++ [α]) { val := i, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
bot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.zero, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
bot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
h': Nat.zero < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.zero, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.zero, isLt := h' }

cons.mk.zero
bot (HList.set (extendBot a (fst, snd)) { val := Nat.zero, isLt := h' } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.zero, isLt := h' } = α) b

Goals accomplished! 🐙
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
bot (HList.set (extendBot a (fst, snd)) { val := Nat.succ n, isLt := h' } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
bot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
bot (HList.set (extendBot a snd) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } b) = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ
cast (_ : List.get (tail ++ [α]) { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) } = α) b = cast (_ : List.get (head :: tail ++ [α]) { val := Nat.succ n, isLt := h' } = α) b
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
¬{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }

cons.mk.succ.h
¬{ val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail
α: Type u_1
a: α
head✝: Type u_1
tail✝: List (Type u_1)
ih: {i : Fin (List.length (List.append tail [α]))} (σ : Assg tail), ¬i.val < List.length tail {b : List.get (List.append tail [α]) i}, bot (HList.set (extendBot a σ) i b) = cast (_ : List.get (tail ++ [α]) i = α) b
i: Fin (List.length (List.append (head :: tail) [α]))
fst✝: head
snd✝: HList tail
n✝: Nat
h': Nat.succ n < List.length (List.append (head :: tail) [α])
h: ¬{ val := Nat.succ n, isLt := h' }.val < List.length (head :: tail)
b: List.get (List.append (head :: tail) [α]) { val := Nat.succ n, isLt := h' }
h'': { val := n, isLt := (_ : Nat.succ n List.length (tail ++ [α])) }.val < List.length tail

cons.mk.succ.h
False

Goals accomplished! 🐙
theorem
eval_S: ∀ {m : Type → Type} {ω : Type} {Γ Δ : List Type} {α : Type} {b c : Bool} {β : Type} {a : α} {ρ : HList Γ} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Stmt m ω Γ (List.append Δ [α]) b c β), StateT.run (Stmt.eval (a :: ρ) (S s) σ) a = do let x ← Stmt.eval ρ s (Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot σ, Assg.bot σ)
eval_S
[
Monad: (Type ?u.552273 → Type ?u.552272) → Type (max (?u.552273 + 1) ?u.552272)
Monad
m: Type → Type
m
] [
LawfulMonad: (m : Type → Type) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type
m
] : (
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
:
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
m: Type → Type
m
ω: Type
ω
Γ: List Type
Γ
(
Δ: List Type
Δ
++ [
α: Type
α
])
b: Bool
b
c: Bool
c
β: Type
β
),
StateT.run: {σ : Type} → {m : Type → Type} → {α : Type} → StateT σ m α → σ → m (α × σ)
StateT.run
((
S: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → [inst : Monad m] → Stmt m ω Γ (List.append Δ [α]) b c β → Stmt (StateT α m) ω (α :: Γ) Δ b c β
S
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
).
eval: {m : Type → Type} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
(
a: α
a
::
ρ: HList Γ
ρ
)
σ: Assg Δ
σ
)
a: α
a
=
s: Stmt m ω Γ (List.append Δ [α]) b c β
s
.
eval: {m : Type → Type} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: HList Γ
ρ
(
Assg.extendBot: {α : Type} → α → {Γ : List Type} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
a: α
a
σ: Assg Δ
σ
) >>= fun |
r: Neut ω β b c
r
::
σ: HList (List.append Δ [α])
σ
=>
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
((
r: Neut ω β b c
r
::
Assg.dropBot: {α : Type} → {Γ : List Type} → Assg (List.append Γ [α]) → Assg Γ
Assg.dropBot
σ: HList (List.append Δ [α])
σ
),
Assg.bot: {α : Type} → {Γ : List Type} → Assg (List.append Γ [α]) → α
Assg.bot
σ: HList (List.append Δ [α])
σ
) :=
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
a: α
ρ: HList Γ
σ: Assg Δ
inst✝¹: Monad m

(s : Stmt m ω Γ (List.append Δ [α]) b c β), StateT.run (Stmt.eval (a :: ρ) (S s) σ) a = do let x Stmt.eval ρ s (Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot σ, Assg.bot σ)
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
a: α
ρ: HList Γ
σ: Assg Δ
inst✝¹: Monad m

{Δ' : List Type} (s : Stmt m ω Γ Δ' b c β) (h : Δ' = List.append Δ [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = Δ') Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
a: α
ρ: HList Γ
σ: Assg Δ
inst✝¹: Monad m
Δ': List Type
s: Stmt m ω Γ Δ' b c β
h: Δ' = List.append Δ [α]

StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = Δ') Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))
m: Type Type
ω: Type
Γ, Δ: List Type
α: Type
b, c: Bool
β: Type
a: α
ρ: HList Γ
σ: Assg Δ
inst✝¹: Monad m
Δ': List Type
s: Stmt m ω Γ Δ' b c β
h: Δ' = List.append Δ [α]

StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = Δ') Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e✝: Assg Γ✝ Assg (List.append Δ [α]) Bool
s₁✝, s₂✝: Stmt m ω Γ✝ (List.append Δ [α]) b c α

ite
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.ite e s₁ s₂)) σ) a = do let x Stmt.eval ρ (Stmt.ite e s₁ s₂) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
s₁: Stmt m ω Γ✝ (List.append Δ [α]) b c α
s₂: Stmt m ω (α :: Γ✝) (List.append Δ [α]) b c β
ih₁: {Δ_1 : List Type} {a : α} {ρ : HList Γ✝} {σ : Assg Δ_1} (h : List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s₁)) σ) a = do let x Stmt.eval ρ s₁ ((_ : List.append Δ_1 [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))
ih₂: {Δ_1 : List Type} {a : α} {ρ : HList (α :: Γ✝)} {σ : Assg Δ_1} (h : List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s₂)) σ) a = do let x Stmt.eval ρ s₂ ((_ : List.append Δ_1 [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))

bind
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.bind s₁ s₂)) σ) a = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
s₁: Stmt m ω Γ✝ (List.append Δ [α]) b c α
s₂: Stmt m ω (α :: Γ✝) (List.append Δ [α]) b c β
ih₂: {Δ_1 : List Type} {a : α} {ρ : HList (α :: Γ✝)} {σ : Assg Δ_1} (h : List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s₂)) σ) a = do let x Stmt.eval ρ s₂ ((_ : List.append Δ_1 [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))
ih₁: (a : α) (ρ : HList Γ✝) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s₁)) σ) a = do let x Stmt.eval ρ s₁ ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

bind
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.bind s₁ s₂)) σ) a = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
s₁: Stmt m ω Γ✝ (List.append Δ [α]) b c α
s₂: Stmt m ω (α :: Γ✝) (List.append Δ [α]) b c β
ih₁: (a : α) (ρ : HList Γ✝) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s₁)) σ) a = do let x Stmt.eval ρ s₁ ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
ih₂: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s₂)) σ) a = do let x Stmt.eval ρ s₂ ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

bind
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.bind s₁ s₂)) σ) a = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) α
s: Stmt m ω Γ✝ :: List.append Δ [α]) b c β
ih: {Δ_1 : List Type} {a : α} {ρ : HList Γ✝} {σ : Assg Δ_1} (h : α :: List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ_1 [α] = α :: List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))

letmut
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.letmut e s)) σ) a = do let x Stmt.eval ρ (Stmt.letmut e s) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) α
s: Stmt m ω Γ✝ :: List.append Δ [α]) b c β
ih: (a : α) (ρ : HList Γ✝) (σ : Assg (α :: Δ)), StateT.run (Stmt.eval (a :: ρ) (S ((_ : α :: List.append Δ [α] = α :: List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append (α :: Δ) [α] = α :: List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : α :: List.append Δ [α] = α :: List.append Δ [α]) σ), Assg.bot ((_ : α :: List.append Δ [α] = α :: List.append Δ [α]) σ))

letmut
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.letmut e s)) σ) a = do let x Stmt.eval ρ (Stmt.letmut e s) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: {Δ_1 : List Type} {a : α} {ρ : HList (α :: Γ✝)} {σ : Assg Δ_1} (h : List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ_1 [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))

sfor
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.sfor e s)) σ) a = do let x Stmt.eval ρ (Stmt.sfor e s) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

sfor
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.sfor e s)) σ) a = do let x Stmt.eval ρ (Stmt.sfor e s) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

sfor
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ (e ρ (Assg.extendBot a σ))) a = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a σ) (e ρ (Assg.extendBot a σ)) pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: {Δ_1 : List Type} {a : α} {ρ : HList (α :: Γ✝)} {σ : Assg Δ_1} (h : List.append Δ [α] = List.append Δ_1 [α]), StateT.run (Stmt.eval (a :: ρ) (S (h s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ_1 [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot (h σ), Assg.bot (h σ))

sfor
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.sfor e s)) σ) a = do let x Stmt.eval ρ (Stmt.sfor e s) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

sfor
StateT.run (Stmt.eval.go (a' :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ (e ρ (Assg.extendBot a' σ))) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) (e ρ (Assg.extendBot a' σ)) pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

a :: ρ
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

a' :: ρ
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

a' :: ρ
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

a :: ρ
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α
h: a = a'

a :: ρ
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
a': α

sfor
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ (e ρ (Assg.extendBot a' σ))) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) (e ρ (Assg.extendBot a' σ)) pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
σ: Assg Δ
a': α

sfor.nil
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ []) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) [] pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
head✝: α
tail✝: List α
σ: Assg Δ
a': α
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ (head :: tail)) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) (head :: tail) pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
σ: Assg Δ
a': α

sfor.nil
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ []) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) [] pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
α✝: Type
b✝, c✝: Bool
Δ: List Type
a: α
ρ: HList Γ✝
e: Assg Γ✝ Assg (List.append Δ [α]) List α
s: Stmt m ω (α :: Γ✝) (List.append Δ [α]) true true Unit
ih: (a : α) (ρ : HList (α :: Γ✝)) (σ : Assg Δ), StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) s)) σ) a = do let x Stmt.eval ρ s ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))
head✝: α
tail✝: List α
σ: Assg Δ
a': α
StateT.run (Stmt.eval.go (a :: ρ) c b α (Stmt.bind (Stmt.expr fun x x => get) (Stmt.mapAssg S.shadowSnd (S s))) σ (head :: tail)) a' = do let x Stmt.eval.go ρ c b α s (Assg.extendBot a' σ) (head :: tail) pure (x.fst :: Assg.dropBot x.snd, Assg.bot x.snd)

Goals accomplished! 🐙
m: Type Type
ω: Type
Γ: List Type
α: Type
b, c: Bool
β: Type
inst✝¹: Monad m
Δ', Γ✝: List Type
b✝, c✝: Bool
α✝: Type
Δ: List Type
a: α
ρ: HList Γ✝
σ: Assg Δ
e✝: Assg Γ✝ Assg (List.append Δ [α]) Bool
s₁✝, s₂✝: Stmt m ω Γ✝ (List.append Δ [α]) b c α

ite
StateT.run (Stmt.eval (a :: ρ) (S ((_ : List.append Δ [α] = List.append Δ [α]) Stmt.ite e s₁ s₂)) σ) a = do let x Stmt.eval ρ (Stmt.ite e s₁ s₂) ((_ : List.append Δ [α] = List.append Δ [α]) Assg.extendBot a σ) match x with | (r, σ) => pure (r :: Assg.dropBot ((_ : List.append Δ [α] = List.append Δ [α]) σ), Assg.bot ((_ : List.append Δ [α] = List.append Δ [α]) σ))

Goals accomplished! 🐙
theorem
HList.eq_nil: ∀ (as : HList ∅), as = ∅
HList.eq_nil
(
as: HList ∅
as
:
HList: List (Type u_1) → Type u_1
HList
: List (Type u_1)
) :
as: HList ∅
as
= :=
rfl: ∀ {α : Type u_1} {a : α}, a = a
rfl
attribute [local simp]
ExceptT.run_bind: ∀ {m : Type u_1 → Type u_2} {ε α β : Type u_1} {f : α → ExceptT ε m β} [inst : Monad m] (x : ExceptT ε m α), ExceptT.run (x >>= f) = do let x ← ExceptT.run x match x with | Except.ok x => ExceptT.run (f x) | Except.error e => pure (Except.error e)
ExceptT.run_bind
theorem
eval_L: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {b c : Bool} {α : Type} {ρ : Assg Γ} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Stmt m ω Γ Δ b c α), Stmt.eval ρ (L s) σ = ExceptT.lift (Stmt.eval ρ s σ)
eval_L
[
Monad: (Type ?u.632180 → Type ?u.632179) → Type (max (?u.632180 + 1) ?u.632179)
Monad
m: Type → Type u_1
m
] [
LawfulMonad: (m : Type → Type u_1) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ b c α
s
:
Stmt: (Type → Type u_1) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 u_1)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
) : (
L: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ b c α
L
s: Stmt m ω Γ Δ b c α
s
).
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
=
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
(
s: Stmt m ω Γ Δ b c α
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
) :=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

Stmt.eval ρ (L s) σ = ExceptT.lift (Stmt.eval ρ s σ)
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (L s) σ) = ExceptT.run (ExceptT.lift (Stmt.eval ρ s σ))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (L s) σ) = ExceptT.run (ExceptT.lift (Stmt.eval ρ s σ))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval ρ (L (Stmt.sfor e s)) σ) = ExceptT.run (ExceptT.lift (Stmt.eval ρ (Stmt.sfor e s) σ))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval.go ρ c b α (L s) σ (e ρ σ)) = ExceptT.run (ExceptT.lift (Stmt.eval.go ρ c b α s σ (e ρ σ)))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c b α (L s) σ []) = ExceptT.run (ExceptT.lift (Stmt.eval.go ρ c b α s σ []))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c b α (L s) σ (head :: tail)) = ExceptT.run (ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c b α (L s) σ []) = ExceptT.run (ExceptT.lift (Stmt.eval.go ρ c b α s σ []))
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c b α (L s) σ (head :: tail)) = ExceptT.run (ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)))

Goals accomplished! 🐙
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
ρ: Assg Γ✝
σ: Assg Δ✝

h.assg
ExceptT.run (Stmt.eval ρ (L (Stmt.assg x e)) σ) = ExceptT.run (ExceptT.lift (Stmt.eval ρ (Stmt.assg x e) σ))

Goals accomplished! 🐙
theorem
eval_B: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {b c : Bool} {α : Type} {ρ : Assg Γ} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Stmt m ω Γ Δ b c α), Stmt.eval ρ (B s) σ = do let x ← ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
eval_B
[
Monad: (Type ?u.666038 → Type ?u.666037) → Type (max (?u.666038 + 1) ?u.666037)
Monad
m: Type → Type u_1
m
] [
LawfulMonad: (m : Type ?u.665935 → Type ?u.665934) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ b c α
s
:
Stmt: (Type → Type ?u.665969) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 ?u.665969)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
b: Bool
b
c: Bool
c
α: Type
α
) : (
B: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ b c α → Stmt (ExceptT Unit m) ω Γ Δ false c α
B
s: Stmt m ω Γ Δ b c α
s
).
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
= (
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
(
s: Stmt m ω Γ Δ b c α
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
) >>= fun
x: Neut ω α b c × Assg Δ
x
=> match (generalizing := false)
x: Neut ω α b c × Assg Δ
x
with | (
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ: Assg Δ
σ
) | (
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) | (
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
,
σ: Assg Δ
σ
) | (
Neut.rbreak: {ω α : Type} → {c : Bool} → Neut ω α true c
Neut.rbreak
, _) =>
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
() :
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
Unit: Type
Unit
m: Type → Type u_1
m
(
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
false: Bool
false
c: Bool
c
×
Assg: List Type → Type
Assg
Δ: List Type
Δ
)) :=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

Stmt.eval ρ (B s) σ = do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (B s) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ b c α

h
ExceptT.run (Stmt.eval ρ (B s) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ s σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval ρ (B (Stmt.sfor e s)) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ (Stmt.sfor e s) σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor
ExceptT.run (Stmt.eval.go ρ c false α (L s) σ (e ρ σ)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (e ρ σ)) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c false α (L s) σ []) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ []) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c false α (L s) σ (head :: tail)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

h.sfor.nil
ExceptT.run (Stmt.eval.go ρ c false α (L s) σ []) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ []) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
ExceptT.run (Stmt.eval.go ρ c false α (L s) σ (head :: tail)) = ExceptT.run do let x ExceptT.lift (Stmt.eval.go ρ c b α s σ (head :: tail)) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()

Goals accomplished! 🐙
m: Type Type u_1
ω: Type
Γ, Δ: List Type
b, c: Bool
α: Type
inst✝¹: Monad m
Γ✝, Δ✝: List Type
b✝, c✝: Bool
x✝: Fin (List.length Δ✝)
e✝: Assg Γ✝ Assg Δ✝ List.get Δ✝ x
ρ: Assg Γ✝
σ: Assg Δ✝

h.assg
ExceptT.run (Stmt.eval ρ (B (Stmt.assg x e)) σ) = ExceptT.run do let x ExceptT.lift (Stmt.eval ρ (Stmt.assg x e) σ) match b, c, x with | b, c, (Neut.ret o, σ) => pure (Neut.ret o, σ) | b, c, (Neut.val a, σ) => pure (Neut.val a, σ) | b, .(true), (Neut.rcont, σ) => pure (Neut.rcont, σ) | .(true), c, (Neut.rbreak, snd) => throw ()

Goals accomplished! 🐙
@[simp] def
throwOnContinue: {m : Type → Type u_1} → {ω α : Type} → {c : Bool} → {Δ : List Type} → [inst : Monad m] → Neut ω α false c × Assg Δ → ExceptT Unit m (Neut ω α false false × Assg Δ)
throwOnContinue
[
Monad: (Type ?u.788677 → Type ?u.788676) → Type (max (?u.788677 + 1) ?u.788676)
Monad
m: Type → Type u_1
m
] : (
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
false: Bool
false
c: Bool
c
×
Assg: List Type → Type
Assg
Δ: List Type
Δ
)
ExceptT: Type → (Type → Type u_1) → Type → Type u_1
ExceptT
Unit: Type
Unit
m: Type → Type u_1
m
(
Neut: Type → Type → Bool → Bool → Type
Neut
ω: Type
ω
α: Type
α
false: Bool
false
false: Bool
false
×
Assg: List Type → Type
Assg
Δ: List Type
Δ
) | (
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.ret: {ω α : Type} → {b c : Bool} → ω → Neut ω α b c
Neut.ret
o: ω
o
,
σ: Assg Δ
σ
) | (
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) =>
pure: {f : Type → Type u_1} → [self : Pure f] → {α : Type} → α → f α
pure
(
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
,
σ: Assg Δ
σ
) | (
Neut.rcont: {ω α : Type} → {b : Bool} → Neut ω α b true
Neut.rcont
, _) =>
throw: {ε : Type} → {m : Type → Type u_1} → [self : MonadExcept ε m] → {α : Type} → ε → m α
throw
(): Unit
()
theorem
eval_C: ∀ {m : Type → Type u_1} {ω : Type} {Γ Δ : List Type} {c : Bool} {α : Type} {ρ : Assg Γ} {σ : Assg Δ} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Stmt m ω Γ Δ false c α), Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
eval_C
[
Monad: (Type ?u.790186 → Type ?u.790185) → Type (max (?u.790186 + 1) ?u.790185)
Monad
m: Type → Type u_1
m
] [
LawfulMonad: (m : Type ?u.790060 → Type ?u.790059) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type u_1
m
] (
s: Stmt m ω Γ Δ false c α
s
:
Stmt: (Type → Type ?u.790118) → Type → List Type → List Type → Bool → Bool → Type → Type (max 1 ?u.790118)
Stmt
m: Type → Type u_1
m
ω: Type
ω
Γ: List Type
Γ
Δ: List Type
Δ
false: Bool
false
c: Bool
c
α: Type
α
) : (
C: {m : Type → Type u_1} → {ω : Type} → {Γ Δ : List Type} → {c : Bool} → {α : Type} → [inst : Monad m] → Stmt m ω Γ Δ false c α → Stmt (ExceptT Unit m) ω Γ Δ false false α
C
s: Stmt m ω Γ Δ false c α
s
).
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
=
ExceptT.lift: {ε : Type} → {m : Type → Type u_1} → [inst : Monad m] → {α : Type} → m α → ExceptT ε m α
ExceptT.lift
(
s: Stmt m ω Γ Δ false c α
s
.
eval: {m : Type → Type u_1} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
σ: Assg Δ
σ
) >>=
throwOnContinue: {m : Type → Type u_1} → {ω α : Type} → {c : Bool} → {Δ : List Type} → [inst : Monad m] → Neut ω α false c × Assg Δ → ExceptT Unit m (Neut ω α false false × Assg Δ)
throwOnContinue
:=
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
s: Stmt m ω Γ Δ false c α

Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m

(s : Stmt m ω Γ Δ false c α), Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m

{b : Bool} (s' : Stmt m ω Γ Δ b c α) (h : b = false), let s := h s'; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
b': Bool
s: Stmt m ω Γ Δ b' c α
h: b' = false

let s := h s; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
ρ: Assg Γ
σ: Assg Δ
inst✝¹: Monad m
b': Bool
s: Stmt m ω Γ Δ b' c α
h: b' = false

let s := h s; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
b✝, c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
s✝: Stmt m ω Γ✝ :: Δ✝) b c β
ρ: Assg Γ✝
σ: Assg Δ✝
h: b = false

letmut
let s := h Stmt.letmut e s✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
ρ: Assg Γ✝
σ: Assg Δ✝
h: b = false

bind
let s := h Stmt.bind s s'✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
c✝: Bool
α✝, β✝: Type
ρ: Assg Γ✝
σ: Assg Δ✝
s✝: Stmt m ω Γ✝ Δ✝ false c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ false c β

bind
let s := (_ : false = false) Stmt.bind s s'✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
b✝, c✝: Bool
α✝, β✝: Type
s✝: Stmt m ω Γ✝ Δ✝ b c α
s'✝: Stmt m ω (α :: Γ✝) Δ✝ b c β
ρ: Assg Γ✝
σ: Assg Δ✝
h: b = false

bind
let s := h Stmt.bind s s'✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
c✝: Bool
α✝: Type
ρ: Assg Γ✝
σ: Assg Δ✝
h: true = false

sbreak
let s := h Stmt.sbreak; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
ρ: Assg Γ✝
σ: Assg Δ✝
s✝: Stmt m ω Γ✝ :: Δ✝) false c β

letmut
let s := (_ : false = false) Stmt.letmut e s✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

sfor
let s := (_ : false = false) Stmt.sfor e s✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

sfor
Stmt.eval.go ρ false false α (L s) σ (e ρ σ) = ExceptT.lift (Stmt.eval.go ρ c false α s σ (e ρ σ)) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

sfor.nil
Stmt.eval.go ρ false false α (L s) σ [] = ExceptT.lift (Stmt.eval.go ρ c false α s σ []) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
Stmt.eval.go ρ false false α (L s) σ (head :: tail) = ExceptT.lift (Stmt.eval.go ρ c false α s σ (head :: tail)) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
σ: Assg Δ✝

sfor.nil
Stmt.eval.go ρ false false α (L s) σ [] = ExceptT.lift (Stmt.eval.go ρ c false α s σ []) >>= throwOnContinue
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
e: Assg Γ✝ Assg Δ✝ List α
s✝: Stmt m ω (α :: Γ✝) Δ✝ true true Unit
ρ: Assg Γ✝
head✝: α
tail✝: List α
σ: Assg Δ✝
Stmt.eval.go ρ false false α (L s) σ (head :: tail) = ExceptT.lift (Stmt.eval.go ρ c false α s σ (head :: tail)) >>= throwOnContinue

Goals accomplished! 🐙
m: Type Type u_1
ω: Type
Γ, Δ: List Type
c: Bool
α: Type
inst✝¹: Monad m
b': Bool
Γ✝, Δ✝: List Type
α✝: Type
c✝: Bool
β✝: Type
e✝: Assg Γ✝ Assg Δ✝ α
ρ: Assg Γ✝
σ: Assg Δ✝
s✝: Stmt m ω Γ✝ :: Δ✝) false c β

letmut
let s := (_ : false = false) Stmt.letmut e s✝; Stmt.eval ρ (C s) σ = ExceptT.lift (Stmt.eval ρ s σ) >>= throwOnContinue

Goals accomplished! 🐙
theorem
D_eq: ∀ (_x : (m : Type → Type) ×' (Γ : List Type) ×' (α : Type) ×' (_ : Assg Γ) ×' (inst : Monad m) ×' (_ : LawfulMonad m) ×' Stmt m Empty Γ ∅ false false α), D _x.2.2.2.2.2.2 _x.2.2.2.1 = do let x ← Stmt.eval _x.2.2.2.1 _x.2.2.2.2.2.2 ∅ match x with | (Neut.val a, snd) => pure a
D_eq
[
Monad: (Type ?u.924916 → Type ?u.924915) → Type (max (?u.924916 + 1) ?u.924915)
Monad
m: Type → Type
m
] [
LawfulMonad: (m : Type ?u.924921 → Type ?u.924920) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type
m
] : (
s: Stmt m Empty Γ ∅ false false α
s
:
Stmt: (Type → Type) → Type → List Type → List Type → Bool → Bool → Type → Type 1
Stmt
m: Type → Type
m
Empty: Type
Empty
Γ: List Type
Γ
: List Type
false: Bool
false
false: Bool
false
α: Type
α
)
D: {m : Type → Type} → {Γ : List Type} → {α : Type} → [inst : Monad m] → Stmt m Empty Γ ∅ false false α → Assg Γ → m α
D
s: Stmt m Empty Γ ∅ false false α
s
ρ: Assg Γ
ρ
=
s: Stmt m Empty Γ ∅ false false α
s
.
eval: {m : Type → Type} → {Γ : List Type} → {ω : Type} → {Δ : List Type} → {b c : Bool} → {α : Type} → [inst : Monad m] → Assg Γ → Stmt m ω Γ Δ b c α → Assg Δ → m (Neut ω α b c × Assg Δ)
eval
ρ: Assg Γ
ρ
>>= fun (
Neut.val: {ω α : Type} → {b c : Bool} → α → Neut ω α b c
Neut.val
a: α
a
, _) =>
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
a: α
a
|
Stmt.expr: {m : Type → Type ?u.926133} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
e: Assg Γ → Assg ∅ → m α
e
=>
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg m α

D (Stmt.expr e) ρ = do let x Stmt.eval ρ (Stmt.expr e) match x with | (Neut.val a, snd) => pure a

Goals accomplished! 🐙
|
Stmt.bind: {m : Type → Type ?u.926201} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
s₁: Stmt m Empty Γ ∅ false false α✝
s₁
s₂: Stmt m Empty (α✝ :: Γ) ∅ false false α
s₂
=>
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
s₁: Stmt m Empty Γ false false α
s₂: Stmt m Empty (α :: Γ) false false α

D (Stmt.bind s₁ s₂) ρ = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
s₁: Stmt m Empty Γ false false α
s₂: Stmt m Empty (α :: Γ) false false α
ih₁: (ρ : Assg Γ), Monad m LawfulMonad m D s₁ ρ = do let x Stmt.eval ρ s₁ match x with | (Neut.val a, snd) => pure a

D (Stmt.bind s₁ s₂) ρ = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
s₁: Stmt m Empty Γ false false α
s₂: Stmt m Empty (α :: Γ) false false α
ih₁: (ρ : Assg Γ), Monad m LawfulMonad m D s₁ ρ = do let x Stmt.eval ρ s₁ match x with | (Neut.val a, snd) => pure a
ih₂: (ρ : Assg (α :: Γ)), Monad m LawfulMonad m D s₂ ρ = do let x Stmt.eval ρ s₂ match x with | (Neut.val a, snd) => pure a

D (Stmt.bind s₁ s₂) ρ = do let x Stmt.eval ρ (Stmt.bind s₁ s₂) match x with | (Neut.val a, snd) => pure a

Goals accomplished! 🐙
|
Stmt.letmut: {m : Type → Type ?u.926267} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
e: Assg Γ → Assg ∅ → α✝
e
s: Stmt m Empty Γ (α✝ :: ∅) false false α
s
=>
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
e: Assg Γ Assg α
s: Stmt m Empty Γ :: ) false false α

D (Stmt.letmut e s) ρ = do let x Stmt.eval ρ (Stmt.letmut e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
e: Assg Γ Assg α
s: Stmt m Empty Γ :: ) false false α
this: Stmt.numExts (S s) < Nat.succ (Stmt.numExts s)

D (Stmt.letmut e s) ρ = do let x Stmt.eval ρ (Stmt.letmut e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
e: Assg Γ Assg α
s: Stmt m Empty Γ :: ) false false α

D (Stmt.letmut e s) ρ = do let x Stmt.eval ρ (Stmt.letmut e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α, α✝: Type
e: Assg Γ Assg α
s: Stmt m Empty Γ :: ) false false α
this: Stmt.numExts (S s) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (S s) a = do let x Stmt.eval a (S s) match x with | (Neut.val a, snd) => pure a

D (Stmt.letmut e s) ρ = do let x Stmt.eval ρ (Stmt.letmut e s) match x with | (Neut.val a, snd) => pure a

Goals accomplished! 🐙
|
Stmt.ite: {m : Type → Type ?u.926338} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
e: Assg Γ → Assg ∅ → Bool
e
s₁: Stmt m Empty Γ ∅ false false α
s₁
s₂: Stmt m Empty Γ ∅ false false α
s₂
=>
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α

D (Stmt.ite e s₁ s₂) ρ = do let x Stmt.eval ρ (Stmt.ite e s₁ s₂) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α

(if e ρ = true then D s₁ ρ else D s₂ ρ) = do let x if e ρ = true then Stmt.eval ρ s₁ else Stmt.eval ρ s₂ match x with | (Neut.val a, snd) => pure a
;
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α

inl
D s₁ ρ = do let x Stmt.eval ρ s₁ match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α
D s₂ ρ = do let x Stmt.eval ρ s₂ match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α

inl
D s₁ ρ = do let x Stmt.eval ρ s₁ match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α: Type
e: Assg Γ Assg Bool
s₁, s₂: Stmt m Empty Γ false false α
D s₂ ρ = do let x Stmt.eval ρ s₂ match x with | (Neut.val a, snd) => pure a

Goals accomplished! 🐙
|
Stmt.ret: {m : Type → Type ?u.926403} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
e: Assg Γ → Assg ∅ → Empty
e
=> nomatch
e: Assg Γ → Assg ∅ → Empty
e
ρ: Assg Γ
ρ
: Assg ∅
|
Stmt.sfor: {m : Type → Type ?u.926500} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
e: Assg Γ → Assg ∅ → List α✝
e
s: Stmt m Empty (α✝ :: Γ) ∅ true true Unit
s
=>
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit

D (Stmt.sfor e s) ρ = do let x Stmt.eval ρ (Stmt.sfor e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)

D (Stmt.sfor e s) ρ = do let x Stmt.eval ρ (Stmt.sfor e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit

D (Stmt.sfor e s) ρ = do let x Stmt.eval ρ (Stmt.sfor e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a

D (Stmt.sfor e s) ρ = do let x Stmt.eval ρ (Stmt.sfor e s) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a

runCatch (forM (e ρ ) fun x => runCatch (D (C (B s)) (x :: ρ))) = do let x Stmt.eval.go ρ false false α s (e ρ ) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a

nil
runCatch (forM [] fun x => runCatch (D (C (B s)) (x :: ρ))) = do let x Stmt.eval.go ρ false false α s [] match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a
head✝: α
tail✝: List α
runCatch (forM (head :: tail) fun x => runCatch (D (C (B s)) (x :: ρ))) = do let x Stmt.eval.go ρ false false α s (head :: tail) match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a

nil
runCatch (forM [] fun x => runCatch (D (C (B s)) (x :: ρ))) = do let x Stmt.eval.go ρ false false α s [] match x with | (Neut.val a, snd) => pure a
m: Type Type
Γ: List Type
ρ: Assg Γ
inst✝¹: Monad m
α✝: Type
e: Assg Γ Assg List α
s: Stmt m Empty (α :: Γ) true true Unit
this: Stmt.numExts (C (B s)) < Nat.succ (Stmt.numExts s)
ih: (a : Assg (α :: Γ)), D (C (B s)) a = do let x Stmt.eval a (C (B s)) match x with | (Neut.val a, snd) => pure a
head✝: α
tail✝: List α
runCatch (forM (head :: tail) fun x => runCatch (D (C (B s)) (x :: ρ))) = do let x Stmt.eval.go ρ false false α s (head :: tail) match x with | (Neut.val a, snd) => pure a

Goals accomplished! 🐙
termination_by _ s => (
s: Stmt m Empty Γ ∅ false false α
s
.
numExts: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → Stmt m ω Γ Δ b c α → Nat
numExts
,
sizeOf: {α : Type 1} → [self : SizeOf α] → α → Nat
sizeOf
s: Stmt m Empty Γ ∅ false false α
s
) decreasing_by

Goals accomplished! 🐙

The equivalence proof cited in the paper follows from the invariants of D and R.

theorem 
Do.trans_eq_eval: ∀ {m : Type → Type} {α : Type} [inst : Monad m] [inst_1 : LawfulMonad m] (s : Do m α), trans s = ⟦s⟧
Do.trans_eq_eval
[
Monad: (Type → Type) → Type 1
Monad
m: Type → Type
m
] [
LawfulMonad: (m : Type ?u.1069260 → Type ?u.1069259) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type
m
] :
s: Do m α
s
:
Do: (Type → Type) → Type → Type 1
Do
m: Type → Type
m
α: Type
α
,
Do.trans: {m : Type → Type} → {α : Type} → [inst : Monad m] → Do m α → m α
Do.trans
s: Do m α
s
=
s: Do m α
s
:=
m: Type Type
α: Type
inst✝¹: Monad m

(s : Do m α), trans s = s

Goals accomplished! 🐙

Partial Evaluation

We define a new term notation simp [...] in e that rewrites the term e using the given simplification theorems.

open Lean in
open Lean.Parser.Tactic in
open Lean.Meta in
open Lean.Elab in
elab "simp" "[" 
simps: Syntax.TSepArray `Lean.Parser.Tactic.simpLemma ","
simps
:
simpLemma: ParserDescr
simpLemma
,* "]" "in"
e: TSyntax `term
e
:term : term => do -- construct goal `⊢ e = ?x` with fresh metavariable `?x`, simplify, solve by reflexivity, -- and return assigned value of `?x` let
e: Expr
e
Term.elabTermAndSynthesize: Syntax → Option Expr → TermElabM Expr
Term.elabTermAndSynthesize
e: TSyntax `term
e
none: {α : Type} → Option α
none
let
x: Expr
x
mkFreshExprMVar: Option Expr → optParam MetavarKind MetavarKind.natural → optParam Name Name.anonymous → MetaM Expr
mkFreshExprMVar
(← inferType e): Expr
(
inferType: Expr → MetaM Expr
inferType
(← inferType e): Expr
e: Expr
e
(← inferType e): Expr
)
let
goal: Expr
goal
mkFreshExprMVar: Option Expr → optParam MetavarKind MetavarKind.natural → optParam Name Name.anonymous → MetaM Expr
mkFreshExprMVar
(← mkEq e x): Expr
(
mkEq: Expr → Expr → MetaM Expr
mkEq
(← mkEq e x): Expr
e: Expr
e
(← mkEq e x): Expr
x: Expr
x
(← mkEq e x): Expr
)
-- disable ζ-reduction to preserve `let`s
Term.runTactic: MVarId → Syntax → TermElabM Unit
Term.runTactic
goal: Expr
goal
.
mvarId!: Expr → MVarId
mvarId!
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
( `(tactic| (
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
simp
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
(
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
config
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
:= { zeta := false }) [$
simps: Syntax.TSepArray `Lean.Parser.Tactic.simpLemma ","
simps
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
:simpLemma,*];
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
rfl
(← `(tactic| (simp (config := { zeta := false }) [$simps:simpLemma,*]; rfl))): TSyntax `tactic
)))
instantiateMVars: {m : Type → Type} → [inst : Monad m] → [inst : MonadMCtx m] → Expr → m Expr
instantiateMVars
x: Expr
x
-- further clean up generated programs attribute [local simp]
Assg.extendBot: {α : Type u_1} → α → {Γ : List (Type u_1)} → Assg Γ → Assg (List.append Γ [α])
Assg.extendBot
cast: {α β : Sort u} → α = β → α → β
cast
attribute [-simp]
StateT.run'_eq: ∀ {m : Type u_1 → Type u_2} {σ α : Type u_1} [inst : Monad m] (x : StateT σ m α) (s : σ), StateT.run' x s = (fun a => a.fst) <$> StateT.run x s
StateT.run'_eq

We can now use this new notation to completely erase the translation functions from an invocation on the example ex2 from For.lean (manually translated to Stmt).

/-
let mut y := init;
for x in xs do' {
  y ← f y x
};
return y
-/

def 
ex2': {m : Type → Type} → {β α : Type} → [inst : Monad m] → (β → α → m β) → β → List α → m β
ex2'
[
Monad: (Type ?u.1085514 → Type ?u.1085513) → Type (max (?u.1085514 + 1) ?u.1085513)
Monad
m: Type → Type
m
] (
f: β → α → m β
f
:
β: Type
β
α: Type
α
m: Type → Type
m
β: Type
β
) (
init: β
init
:
β: Type
β
) (
xs: List α
xs
:
List: Type → Type
List
α: Type
α
) :
m: Type → Type
m
β: Type
β
:= simp [Do.trans] in
Do.trans: {m : Type → Type} → {α : Type} → [inst : Monad m] → Do m α → m α
Do.trans
(
Stmt.letmut: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → {β : Type} → (Assg Γ → Assg Δ → α) → Stmt m ω Γ (α :: Δ) b c β → Stmt m ω Γ Δ b c β
Stmt.letmut
(fun _ _ =>
init: β
init
) <|
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.sfor: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
(fun _ _ =>
xs: List α
xs
) <| -- `y ← f y x` unfolded to `let z ← f y x; y := z` (A4)
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun ([
x: α
x
]) ([
y: β
y
]) =>
f: β → α → m β
f
y: β
y
x: α
x
)) (
Stmt.assg: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → (x : Fin (List.length Δ)) → (Assg Γ → Assg Δ → List.get Δ x) → Stmt m ω Γ Δ b c Unit
Stmt.assg
0: Nat
0
,
m: Type Type
β, α: Type
inst✝: Monad m
f: β α m β
init: β
xs: List α

0 < List.length (β :: )

Goals accomplished! 🐙
⟩ (fun ([
z: β
z
,
Warning: unused variable `x`
]) _ =>
z: β
z
))) <|
Stmt.ret: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
(fun _ ([
y: β
y
]) =>
y: β
y
))

Compare the output of the two versions - the structure is identical except for unused monadic layers in the formal translation, which would be harder to avoid in this typed approach.

def ex2.{u_1} : {m : Type Type u_1} {α : Type} [inst : Monad m] {β : Type} α m β) β List α m β := fun {m} {α} [Monad m] {β} f init xs => ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
ex2: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → {β : Type} → (β → α → m β) → β → List α → m β
ex2
def ex2' : {m : Type Type} {β α : Type} [inst : Monad m] α m β) β List α m β := fun {m} {β α} [Monad m] f init xs => runCatch (let x := init; StateT.run' (do runCatch (forM xs fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x))) let x get StateT.lift (throw x)) x)
ex2': {m : Type → Type} → {β α : Type} → [inst : Monad m] → (β → α → m β) → β → List α → m β
ex2'

We can evaluate the generated program like any other Lean program, and prove that both versions are equivalent.

3
ex2': {m : Type → Type} → {β α : Type} → [inst : Monad m] → (β → α → m β) → β → List α → m β
ex2'
(m :=
Id: Type → Type
Id
) (fun
a: Nat
a
b: Nat
b
=>
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
(
a: Nat
a
+
b: Nat
b
))
0: Nat
0
[
1: Nat
1
,
2: Nat
2
]
example: ∀ {m : Type → Type} {β α : Type} {init : β} {xs : List α} [inst : Monad m] [inst_1 : LawfulMonad m] (f : β → α → m β), ex2' f init xs = ex2 f init xs
example
[
Monad: (Type → Type) → Type 1
Monad
m: Type → Type
m
] [
LawfulMonad: (m : Type → Type) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type
m
] (
f: β → α → m β
f
:
β: Type
β
α: Type
α
m: Type → Type
m
β: Type
β
) :
ex2': {m : Type → Type} → {β α : Type} → [inst : Monad m] → (β → α → m β) → β → List α → m β
ex2'
f: β → α → m β
f
init: β
init
xs: List α
xs
=
ex2: {m : Type → Type} → {α : Type} → [inst : Monad m] → {β : Type} → (β → α → m β) → β → List α → m β
ex2
f: β → α → m β
f
init: β
init
xs: List α
xs
:=
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

ex2' f init xs = ex2 f init xs
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

ex2' f init xs = ex2 f init xs
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

ex2' f init xs = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

runCatch (let x := init; StateT.run' (do runCatch (forM xs fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x))) let x get StateT.lift (throw x)) x) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

runCatch (let x := init; StateT.run' (do runCatch (forM xs fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x))) let x get StateT.lift (throw x)) x) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
;
m: Type Type
β, α: Type
init: β
xs: List α
inst✝¹: Monad m
f: β α m β

(do let x ExceptT.run (let x := init; StateT.run' (do do let x ExceptT.run (forM xs fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x)) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e let x get StateT.lift (throw x)) x) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM xs fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
;
m: Type Type
β, α: Type
inst✝¹: Monad m
f: β α m β
init: β

nil
(do let x ExceptT.run (let x := init; StateT.run' (do do let x ExceptT.run (forM [] fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x)) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e let x get StateT.lift (throw x)) x) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM [] fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
m: Type Type
β, α: Type
inst✝¹: Monad m
f: β α m β
head✝: α
tail✝: List α
init: β
(do let x ExceptT.run (let x := init; StateT.run' (do do let x ExceptT.run (forM (head :: tail) fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x)) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e let x get StateT.lift (throw x)) x) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM (head :: tail) fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
m: Type Type
β, α: Type
inst✝¹: Monad m
f: β α m β
init: β

nil
(do let x ExceptT.run (let x := init; StateT.run' (do do let x ExceptT.run (forM [] fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x)) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e let x get StateT.lift (throw x)) x) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM [] fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)
m: Type Type
β, α: Type
inst✝¹: Monad m
f: β α m β
head✝: α
tail✝: List α
init: β
(do let x ExceptT.run (let x := init; StateT.run' (do do let x ExceptT.run (forM (head :: tail) fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift get) let x ExceptT.lift (ExceptT.lift (StateT.lift (ExceptT.lift (f x_1 x)))) let _ ExceptT.lift (ExceptT.lift get) ExceptT.lift (ExceptT.lift (set x)) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e let x get StateT.lift (throw x)) x) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch (let y := init; StateT.run' (do forM (head :: tail) fun x => do let y get let y StateT.lift (ExceptCpsT.lift (f y x)) let _ get set y let y get StateT.lift (throw y)) y)

Goals accomplished! 🐙

For one more example, consider ex3 from For.lean.

/-
for xs in xss do' {
  for x in xs do' {
    let b ← p x;
    if b then {
      return some x
    }
  }
};
pure none
-/

def 
ex3': {m : Type → Type} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3'
[
Monad: (Type → Type) → Type 1
Monad
m: Type → Type
m
] (
p: α → m Bool
p
:
α: Type
α
m: Type → Type
m
Bool: Type
Bool
) (
xss: List (List α)
xss
:
List: Type → Type
List
(
List: Type → Type
List
α: Type
α
)) :
m: Type → Type
m
(
Option: Type → Type
Option
α: Type
α
) := simp [Do.trans] in
Do.trans: {m : Type → Type} → {α : Type} → [inst : Monad m] → Do m α → m α
Do.trans
(
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.sfor: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
(fun _ _ =>
xss: List (List α)
xss
) <|
Stmt.sfor: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → List α) → Stmt m ω (α :: Γ) Δ true true Unit → Stmt m ω Γ Δ b c Unit
Stmt.sfor
(fun ([
xs: List α
xs
]) _ =>
xs: List α
xs
) <|
Stmt.bind: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α β : Type} → Stmt m ω Γ Δ b c α → Stmt m ω (α :: Γ) Δ b c β → Stmt m ω Γ Δ b c β
Stmt.bind
(
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun ([
x: α
x
,
Warning: unused variable `xs`
]) _ =>
p: α → m Bool
p
x: α
x
)) (
Stmt.ite: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → Bool) → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α → Stmt m ω Γ Δ b c α
Stmt.ite
(fun ([
b: Bool
b
,
Warning: unused variable `x`
,
Warning: unused variable `xs`
]) _ =>
b: Bool
b
) (
Stmt.ret: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {b c : Bool} → {α : Type} → (Assg Γ → Assg Δ → ω) → Stmt m ω Γ Δ b c α
Stmt.ret
(fun ([
Warning: unused variable `b`
,
x: α
x
,
Warning: unused variable `xs`
]) _ =>
some: {α : Type} → α → Option α
some
x: α
x
)) (
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
(): Unit
()
)))) (
Stmt.expr: {m : Type → Type} → {ω : Type} → {Γ Δ : List Type} → {α : Type} → {b c : Bool} → (Assg Γ → Assg Δ → m α) → Stmt m ω Γ Δ b c α
Stmt.expr
(fun _ _ =>
pure: {f : Type → Type} → [self : Pure f] → {α : Type} → α → f α
pure
none: {α : Type} → Option α
none
)))
def ex3.{u_1} : {m : Type Type u_1} {α : Type} [inst : Monad m] m Bool) List (List α) m (Option α) := fun {m} {α} [Monad m] p xss => ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
ex3: {m : Type → Type u_1} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3
def ex3' : {m : Type Type} {α : Type} [inst : Monad m] m Bool) List (List α) m (Option α) := fun {m} {α} [Monad m] p xss => runCatch do runCatch (forM xss fun x => runCatch (runCatch (forM x fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ())))))))) ExceptT.lift (pure none)
ex3': {m : Type → Type} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3'
some 2
ex3': {m : Type → Type} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3'
(m :=
Id: Type → Type
Id
) (fun
n: Nat
n
=>
n: Nat
n
%
2: Nat
2
==
0: Nat
0
) [[
1: Nat
1
,
3: Nat
3
], [
2: Nat
2
,
4: Nat
4
]]
example: ∀ {m : Type → Type} {α : Type} [inst : Monad m] [inst_1 : LawfulMonad m] (p : α → m Bool) (xss : List (List α)), ex3' p xss = ex3 p xss
example
[
Monad: (Type → Type) → Type 1
Monad
m: Type → Type
m
] [
LawfulMonad: (m : Type ?u.1130706 → Type ?u.1130705) → [inst : Monad m] → Prop
LawfulMonad
m: Type → Type
m
] (
p: α → m Bool
p
:
α: Type
α
m: Type → Type
m
Bool: Type
Bool
) (
xss: List (List α)
xss
:
List: Type → Type
List
(
List: Type → Type
List
α: Type
α
)) :
ex3': {m : Type → Type} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3'
p: α → m Bool
p
xss: List (List α)
xss
=
ex3: {m : Type → Type} → {α : Type} → [inst : Monad m] → (α → m Bool) → List (List α) → m (Option α)
ex3
p: α → m Bool
p
xss: List (List α)
xss
:=
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

ex3' p xss = ex3 p xss
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

ex3' p xss = ex3 p xss
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

ex3' p xss = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

(runCatch do runCatch (forM xss fun x => runCatch (runCatch (forM x fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ())))))))) ExceptT.lift (pure none)) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

(runCatch do runCatch (forM xss fun x => runCatch (runCatch (forM x fun x => runCatch do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ())))))))) ExceptT.lift (pure none)) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

(do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)

(do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool

nil
(do let x ExceptT.run do do let x ExceptT.run (forM [] fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM [] fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

Goals accomplished! 🐙
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xs: List α
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons
(do let x ExceptT.run do do let x ExceptT.run (forM (xs :: xss) fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM (xs :: xss) fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xs: List α
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons
(do let x ExceptT.run (ExceptT.run (ExceptT.run (ExceptT.run (forM xs fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)))) let x match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => ExceptT.run (forM xss fun x => do let x ExceptT.run (ExceptT.run (forM x fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)) let x match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => pure (Except.ok none) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) = ExceptCpsT.runCatch do forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.nil
(do let x ExceptT.run (ExceptT.run (ExceptT.run (ExceptT.run (forM [] fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)))) let x match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => ExceptT.run (forM xss fun x => do let x ExceptT.run (ExceptT.run (forM x fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)) let x match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => pure (Except.ok none) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) = ExceptCpsT.runCatch do forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
head✝: α
tail✝: List α
(do let x ExceptT.run (ExceptT.run (ExceptT.run (ExceptT.run (forM (head :: tail) fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)))) let x match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => ExceptT.run (forM xss fun x => do let x ExceptT.run (ExceptT.run (forM x fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)) let x match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => pure (Except.ok none) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) = ExceptCpsT.runCatch do forM (head :: tail) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

cons.nil
(do let x ExceptT.run (ExceptT.run (ExceptT.run (ExceptT.run (forM [] fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)))) let x match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => ExceptT.run (forM xss fun x => do let x ExceptT.run (ExceptT.run (forM x fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)) let x match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => pure (Except.ok none) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) = ExceptCpsT.runCatch do forM [] fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
m: Type Type
α: Type
inst✝¹: Monad m
p: α m Bool
xss: List (List α)
ih: (do let x ExceptT.run do do let x ExceptT.run (forM xss fun x => do let x ExceptT.run do let x ExceptT.run (forM x fun x => do let x ExceptT.run do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x))))) if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (pure ()))))) match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e match x with | Except.ok x => pure x | Except.error e => pure e) match x with | Except.ok x => pure x | Except.error e => pure e ExceptT.lift (pure none) match x with | Except.ok x => pure x | Except.error e => pure e) = ExceptCpsT.runCatch do forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)
head✝: α
tail✝: List α
(do let x ExceptT.run (ExceptT.run (ExceptT.run (ExceptT.run (forM (head :: tail) fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)))) let x match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => ExceptT.run (forM xss fun x => do let x ExceptT.run (ExceptT.run (forM x fun x => do let x_1 ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (p x)))) let x ExceptT.run (if x_1 = true then ExceptT.lift (ExceptT.lift (ExceptT.lift (ExceptT.lift (throw (some x))))) else pure ()) match x with | Except.ok x => pure x | Except.error x => pure x)) let x match x with | Except.ok x => ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e)) match x with | Except.ok x => do let x ExceptT.run (match x with | Except.ok x => pure x | Except.error x => pure x) match x with | Except.ok x => pure (Except.ok none) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) | Except.error e => pure (Except.error e) match x with | Except.ok x => pure x | Except.error x => pure x) = ExceptCpsT.runCatch do forM (head :: tail) fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) forM xss fun xs => forM xs fun x => do let b ExceptCpsT.lift (p x) if b = true then throw (some x) else ExceptCpsT.lift (pure ()) ExceptCpsT.lift (pure none)

Goals accomplished! 🐙

While it would be possible to override our do' notation such that its named syntax is first translated to nameless Stmt constructors and then applied to simp [Do.trans] in, for demonstration purposes we decided to encode these examples manually. In practice, the macro implementation remains more desirable as mentioned in the paper.