├── .gitignore ├── lean-toolchain ├── Lean4Lean ├── Verify.lean ├── Std │ ├── Control.lean │ ├── Variable!.lean │ ├── NodupKeys.lean │ ├── PersistentHashMap.lean │ ├── ToExpr.lean │ ├── SMap.lean │ ├── HashMap.lean │ └── Basic.lean ├── Theory.lean ├── List.lean ├── Stream.lean ├── Theory │ ├── Inductive.lean │ ├── Typing │ │ ├── InductiveLemmas.lean │ │ ├── EnvLemmas.lean │ │ ├── Env.lean │ │ ├── QuotLemmas.lean │ │ ├── Meta.lean │ │ ├── Basic.lean │ │ └── UniqueTyping.lean │ ├── Quot.lean │ ├── VDecl.lean │ ├── VEnv.lean │ ├── Meta.lean │ └── VLevel.lean ├── Declaration.lean ├── Instantiate.lean ├── LocalContext.lean ├── PtrEq.lean ├── Verify │ ├── NameGenerator.lean │ ├── VLCtx.lean │ ├── TypeChecker.lean │ ├── Environment │ │ ├── Basic.lean │ │ └── Lemmas.lean │ ├── Level.lean │ ├── Typing │ │ ├── ConditionallyTyped.lean │ │ └── Expr.lean │ ├── TypeChecker │ │ ├── Reduce.lean │ │ └── WHNF.lean │ └── LocalContext.lean ├── ForEachExprV.lean ├── EquivManager.lean ├── Inductive │ └── Reduce.lean ├── Expr.lean ├── Quot.lean ├── Environment.lean ├── Environment │ └── Basic.lean ├── Experimental │ ├── ParallelReduction.lean │ └── StratifiedUntyped.lean └── Level.lean ├── bugs-found.md ├── lake-manifest.json ├── lakefile.toml ├── Lean4Lean.lean ├── divergences.md ├── README.md └── LICENSE /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.23.0 2 | -------------------------------------------------------------------------------- /Lean4Lean/Verify.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.Typing.Lemmas 2 | -------------------------------------------------------------------------------- /Lean4Lean/Std/Control.lean: -------------------------------------------------------------------------------- 1 | 2 | instance [Alternative m] : MonadLift Option m := ⟨fun | none => failure | some a => pure a⟩ 3 | -------------------------------------------------------------------------------- /Lean4Lean/Theory.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.EnvLemmas 2 | import Lean4Lean.Theory.Typing.Strong 3 | import Lean4Lean.Theory.Typing.UniqueTyping 4 | -------------------------------------------------------------------------------- /Lean4Lean/List.lean: -------------------------------------------------------------------------------- 1 | 2 | @[specialize] def List.all2 (R : α → α → Bool) : List α → List α → Bool 3 | | a1 :: as1, a2 :: as2 => R a1 a2 && all2 R as1 as2 4 | | [], [] => true 5 | | _, _ => false 6 | -------------------------------------------------------------------------------- /Lean4Lean/Stream.lean: -------------------------------------------------------------------------------- 1 | 2 | structure NatStream where i := 0 deriving Inhabited 3 | 4 | instance : ToStream NatStream NatStream := ⟨id⟩ 5 | instance : Stream NatStream Nat := ⟨fun ⟨r⟩ => some (r, ⟨r + 1⟩)⟩ 6 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Inductive.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VDecl 2 | 3 | namespace Lean4Lean 4 | 5 | def VInductDecl.WF (env : VEnv) (decl : VInductDecl) : Prop := sorry 6 | 7 | def VEnv.addInduct (env : VEnv) (decl : VInductDecl) : Option VEnv := sorry 8 | -------------------------------------------------------------------------------- /bugs-found.md: -------------------------------------------------------------------------------- 1 | # Bugs found in the kernel as a result of formalization 2 | 3 | * https://leanprover.zulipchat.com/#narrow/channel/270676-lean4/topic/Soundness.20bug.3A.20hasLooseBVars.20is.20not.20conservative/near/521286338 4 | * https://github.com/leanprover/lean4/issues/10475 5 | * https://github.com/leanprover/lean4/issues/10511 6 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/InductiveLemmas.lean: -------------------------------------------------------------------------------- 1 | import Std 2 | import Lean4Lean.Theory.Typing.Lemmas 3 | import Lean4Lean.Theory.Typing.Env 4 | 5 | namespace Lean4Lean 6 | namespace VEnv 7 | 8 | theorem addInduct_WF (henv : Ordered env) (hdecl : decl.WF env) 9 | (henv' : addInduct env decl = some env') : Ordered env' := 10 | sorry 11 | -------------------------------------------------------------------------------- /Lean4Lean/Declaration.lean: -------------------------------------------------------------------------------- 1 | import Lean.Declaration 2 | 3 | 4 | namespace Lean 5 | namespace ReducibilityHints 6 | 7 | def lt' : ReducibilityHints → ReducibilityHints → Bool -- lean4#2750 8 | | _, .opaque => false 9 | | .abbrev, _ => false 10 | | .opaque, _ => true 11 | | _, .abbrev => true 12 | | .regular d₁, .regular d₂ => d₁ < d₂ 13 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/batteries", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "d117e2c28cba42e974bc22568ac999492a34e812", 9 | "name": "batteries", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.23.0", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "lean4lean", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "lean4lean" 2 | defaultTargets = ["Lean4Lean", "lean4lean", "Lean4Lean.Theory", "Lean4Lean.Verify"] 3 | 4 | [[require]] 5 | name = "batteries" 6 | git = "https://github.com/leanprover-community/batteries" 7 | rev = "v4.23.0" 8 | 9 | [[lean_lib]] 10 | name = "Lean4Lean" 11 | 12 | [[lean_lib]] 13 | name = "Lean4Lean.Theory" 14 | globs = ["Lean4Lean.Theory.*"] 15 | 16 | [[lean_lib]] 17 | name = "Lean4Lean.Verify" 18 | globs = ["Lean4Lean.Verify.*"] 19 | 20 | [[lean_lib]] 21 | name = "Lean4Lean.Experimental" 22 | globs = ["Lean4Lean.Experimental.*"] 23 | 24 | [[lean_exe]] 25 | name = "lean4lean" 26 | root = "Main" 27 | supportInterpreter = true 28 | -------------------------------------------------------------------------------- /Lean4Lean/Std/Variable!.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | macro "variable!" args:(ppSpace colGt bracketedBinder)+ " in" 4 | cmd:ppDedent(command) : command => do 5 | let mut incls := #[] 6 | for a in args do 7 | match a with 8 | | `(bracketedBinder| ($as* $[: $_]?)) 9 | | `(bracketedBinder| {$as* $[: $_]?}) 10 | | `(bracketedBinder| ⦃$as* $[: $_]?⦄) => 11 | for a in as do 12 | match a with | `(_) => pure () | _ => incls := incls.push ⟨a.raw⟩ 13 | | `(bracketedBinder| [$a : $_]) => 14 | match a with | `(_) => pure () | _ => incls := incls.push ⟨a.raw⟩ 15 | | _ => pure () 16 | `(command| variable $args* in include $incls* in $(⟨cmd.raw⟩)) 17 | -------------------------------------------------------------------------------- /Lean4Lean/Instantiate.lean: -------------------------------------------------------------------------------- 1 | import Lean.Expr 2 | import Lean.LocalContext 3 | import Lean.Util.InstantiateLevelParams 4 | 5 | namespace Lean 6 | namespace Expr 7 | 8 | def cheapBetaReduce (e : Expr) : Expr := Id.run do 9 | if !e.isApp then return e 10 | let fn := e.getAppFn 11 | if !fn.isLambda then return e 12 | let args := e.getAppArgs 13 | let rec cont i fn := 14 | if !fn.hasLooseBVars then 15 | mkAppRange fn i args.size args 16 | else if let .bvar n := fn then 17 | assert! n < i 18 | mkAppRange args[i - n - 1]! i args.size args 19 | else 20 | e 21 | let rec loop i fn := 22 | if i < args.size then 23 | match fn with 24 | | .lam _ _ body .. => loop (i + 1) body 25 | | _ => cont i fn 26 | else cont i fn 27 | return loop 0 fn 28 | 29 | end Expr 30 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Quot.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VEnv 2 | import Lean4Lean.Theory.Meta 3 | 4 | namespace Lean4Lean 5 | 6 | def eqConst := vconst(type_of% @Eq) 7 | def quotConst := vconst(type_of% @Quot) 8 | def quotMkConst := vconst(type_of% @Quot.mk) 9 | def quotLiftConst := vconst(type_of% @Quot.lift) 10 | def quotIndConst := vconst(type_of% @Quot.ind) 11 | def quotDefEq := vdefeq(α r β f c a => @Quot.lift α r β f c (Quot.mk r a) ≡ f a) 12 | 13 | def VEnv.QuotReady (env : VEnv) : Prop := 14 | env.constants ``Eq = some eqConst 15 | 16 | def VEnv.addQuot (env : VEnv) : Option VEnv := do 17 | let env ← env.addConst ``Quot quotConst 18 | let env ← env.addConst ``Quot.mk quotMkConst 19 | let env ← env.addConst ``Quot.lift quotLiftConst 20 | let env ← env.addConst ``Quot.ind quotIndConst 21 | env.addDefEq quotDefEq 22 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/VDecl.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VEnv 2 | 3 | namespace Lean4Lean 4 | 5 | structure VConstVal extends VConstant where 6 | name : Name 7 | 8 | structure VDefVal extends VConstVal where 9 | value : VExpr 10 | 11 | def VDefVal.toDefEq (v : VDefVal) : VDefEq := 12 | ⟨v.uvars, .const v.name (VLevel.params v.uvars), v.value, v.type⟩ 13 | 14 | structure VInductiveType extends VConstVal where 15 | ctors : List VConstVal 16 | 17 | structure VInductDecl where 18 | uvars : Nat 19 | nparams : Nat 20 | types : List VInductiveType 21 | 22 | inductive VDecl where 23 | /-- Reserve a constant name, which cannot be used in expressions. 24 | Used to represent unsafe declarations in safe mode -/ 25 | | block (n : Name) 26 | | axiom (_ : VConstVal) 27 | | def (_ : VDefVal) 28 | | opaque (_ : VDefVal) 29 | | example (_ : VDefVal) 30 | | quot 31 | | induct (_ : VInductDecl) 32 | -------------------------------------------------------------------------------- /Lean4Lean/LocalContext.lean: -------------------------------------------------------------------------------- 1 | import Lean.LocalContext 2 | 3 | namespace Lean4Lean 4 | open Lean 5 | 6 | class MonadLocalNameGenerator (m : Type u → Type v) where 7 | withFreshId : (Name → m α) → m α 8 | export MonadLocalNameGenerator (withFreshId) 9 | 10 | instance [MonadLocalNameGenerator m] : MonadLocalNameGenerator (ReaderT ρ m) where 11 | withFreshId f c := withFreshId (f · c) 12 | 13 | @[inline] def withLocalDecl [Monad m] [MonadLocalNameGenerator m] [MonadWithReaderOf LocalContext m] 14 | (name : Name) (ty : Expr) (bi : BinderInfo) (k : Expr → m α) : m α := 15 | withFreshId fun id => do 16 | withReader (·.mkLocalDecl ⟨id⟩ name ty bi) <| k <| .fvar ⟨id⟩ 17 | 18 | @[inline] def withLetDecl [Monad m] [MonadLocalNameGenerator m] [MonadWithReaderOf LocalContext m] 19 | (name : Name) (ty val : Expr) (k : Expr → m α) : m α := 20 | withFreshId fun id => do 21 | withReader (·.mkLetDecl ⟨id⟩ name ty val) <| k <| .fvar ⟨id⟩ 22 | -------------------------------------------------------------------------------- /Lean4Lean/PtrEq.lean: -------------------------------------------------------------------------------- 1 | import Lean.Declaration 2 | 3 | namespace Lean4Lean 4 | open Lean 5 | 6 | /-- 7 | Pointer equality is not a safe function, but the kernel uses it anyway for some fast paths. 8 | We cannot use `withPtrEq` because it is not a mere optimization before a true equality test - 9 | the kernel actually has different behavior (will reject some inputs that it would not otherwise) 10 | if identical objects spontaneously get different addresses. 11 | 12 | We use this function to try to isolate uses of pointer equality in the kernel, 13 | and type-restrict it to avoid thorny questions about equality of closures or the like. 14 | -/ 15 | opaque ptrEqExpr (a b : Expr) : Bool := unsafe ptrAddrUnsafe a == ptrAddrUnsafe b 16 | 17 | axiom ptrEqExpr_eq : ptrEqExpr a b → a = b 18 | 19 | /-- See `ptrEqExpr`. -/ 20 | opaque ptrEqConstantInfo (a b : ConstantInfo) : Bool := unsafe ptrAddrUnsafe a == ptrAddrUnsafe b 21 | 22 | axiom ptrEqConstantInfo_eq : ptrEqConstantInfo a b → a = b 23 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/EnvLemmas.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Lemmas 2 | import Lean4Lean.Theory.Typing.Env 3 | import Lean4Lean.Theory.Typing.QuotLemmas 4 | import Lean4Lean.Theory.Typing.InductiveLemmas 5 | 6 | namespace Lean4Lean 7 | 8 | theorem VEnv.WF.ordered : WF env → Ordered env 9 | | ⟨ds, H⟩ => by 10 | induction H with 11 | | empty => exact .empty 12 | | decl h _ ih => 13 | cases h with 14 | | block h => exact .const ih (by rintro _ ⟨⟩) h 15 | | «axiom» h1 h2 => exact .const ih (by rintro _ ⟨⟩; exact h1) h2 16 | | @«def» env env' ci h1 h2 => 17 | refine .defeq (.const ih (by rintro _ ⟨⟩; exact h1.isType ih ⟨⟩) h2) ⟨?_, ?_⟩ 18 | · simp [VDefVal.toDefEq] 19 | rw [← (h1.levelWF ⟨⟩).2.2.instL_id] 20 | exact .const (addConst_self h2) VLevel.id_WF (by simp) 21 | · exact h1.mono (addConst_le h2) 22 | | «opaque» h1 h2 => exact .const ih (by rintro _ ⟨⟩; exact h1.isType ih ⟨⟩) h2 23 | | «example» _ => exact ih 24 | | quot h1 h2 => exact addQuot_WF ih h1 h2 25 | | induct h1 h2 => exact addInduct_WF ih h1 h2 26 | 27 | instance : CoeOut (VEnv.WF env) env.Ordered := ⟨(·.ordered)⟩ 28 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/NameGenerator.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Expr 2 | import Lean4Lean.Theory.VExpr 3 | 4 | namespace Lean.NameGenerator 5 | 6 | def Reserves (ngen : NameGenerator) (fv : FVarId) : Prop := 7 | ∀ i, fv = ⟨.num ngen.namePrefix i⟩ → i < ngen.idx 8 | 9 | theorem next_reserves_self {ngen : NameGenerator} : 10 | ngen.next.Reserves ⟨ngen.curr⟩ := by rintro _ ⟨⟩; exact Nat.lt_succ_self _ 11 | 12 | theorem not_reserves_self {ngen : NameGenerator} : 13 | ¬ngen.Reserves ⟨ngen.curr⟩ := fun h => Nat.lt_irrefl _ (h _ rfl) 14 | 15 | protected inductive LE : NameGenerator → NameGenerator → Prop where 16 | | le : i ≤ j → NameGenerator.LE ⟨pfx, i⟩ ⟨pfx, j⟩ 17 | 18 | instance : LE NameGenerator := ⟨NameGenerator.LE⟩ 19 | 20 | theorem LE.rfl {ngen : NameGenerator} : ngen ≤ ngen := ⟨Nat.le_refl _⟩ 21 | 22 | theorem LE.trans {ngen₁ ngen₂ ngen₃ : NameGenerator} : ngen₁ ≤ ngen₂ → ngen₂ ≤ ngen₃ → ngen₁ ≤ ngen₃ 23 | | ⟨h₁⟩, ⟨h₂⟩ => ⟨Nat.le_trans h₁ h₂⟩ 24 | 25 | theorem Reserves.mono : ngen ≤ ngen' → Reserves ngen fv → Reserves ngen' fv 26 | | ⟨h₁⟩ => fun H _ hi => Nat.lt_of_lt_of_le (H _ hi) h₁ 27 | 28 | theorem LE.next {ngen : NameGenerator} : ngen ≤ ngen.next := ⟨Nat.le_succ _⟩ 29 | -------------------------------------------------------------------------------- /Lean4Lean/ForEachExprV.lean: -------------------------------------------------------------------------------- 1 | import Lean.Expr 2 | import Lean.Util.MonadCache 3 | 4 | /-! 5 | This is the same as `Expr.forEach` but it uses `StateT` instead of `StateRefT` to avoid opaques. 6 | -/ 7 | namespace Lean 8 | variable {ω : Type} {m : Type → Type} [Monad m] 9 | 10 | namespace ForEachExprV 11 | def visit (g : Expr → m Bool) (e : Expr) : MonadStateCacheT Expr Unit m Unit := 12 | checkCache e fun _ => do 13 | if (← g e) then 14 | match e with 15 | | Expr.forallE _ d b _ => do visit g d; visit g b 16 | | Expr.lam _ d b _ => do visit g d; visit g b 17 | | Expr.letE _ t v b _ => do visit g t; visit g v; visit g b 18 | | Expr.app f a => do visit g f; visit g a 19 | | Expr.mdata _ b => visit g b 20 | | Expr.proj _ _ b => visit g b 21 | | _ => pure () 22 | 23 | end ForEachExprV 24 | 25 | /-- Apply `f` to each sub-expression of `e`. If `f t` returns false, then t's children are not visited. -/ 26 | @[inline] def Expr.forEachV' (e : Expr) (f : Expr → m Bool) : m Unit := 27 | (ForEachExprV.visit f e).run 28 | 29 | @[inline] def Expr.forEachV (e : Expr) (f : Expr → m Unit) : m Unit := 30 | e.forEachV' fun e => do f e; pure true 31 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/VEnv.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VExpr 2 | 3 | namespace Lean4Lean 4 | 5 | structure VConstant where 6 | uvars : Nat 7 | type : VExpr 8 | 9 | structure VDefEq where 10 | uvars : Nat 11 | lhs : VExpr 12 | rhs : VExpr 13 | type : VExpr 14 | 15 | @[ext] structure VEnv where 16 | constants : Name → Option (Option VConstant) 17 | defeqs : VDefEq → Prop 18 | 19 | def VEnv.empty : VEnv where 20 | constants _ := none 21 | defeqs _ := False 22 | 23 | instance : EmptyCollection VEnv := ⟨.empty⟩ 24 | 25 | def VEnv.contains (env : VEnv) (name : Name) := ∃ ci, env.constants name = some (some ci) 26 | 27 | def VEnv.addConst (env : VEnv) (name : Name) (ci : Option VConstant) : Option VEnv := 28 | match env.constants name with 29 | | some _ => none 30 | | none => some { env with constants := fun n => if name = n then some ci else env.constants n } 31 | 32 | def VEnv.addDefEq (env : VEnv) (df : VDefEq) : VEnv := 33 | { env with defeqs := fun x => x = df ∨ env.defeqs x } 34 | 35 | structure VEnv.LE (env1 env2 : VEnv) : Prop where 36 | constants : env1.constants n = some a → env2.constants n = some a 37 | defeqs : env1.defeqs df → env2.defeqs df 38 | 39 | instance : LE VEnv := ⟨VEnv.LE⟩ 40 | 41 | theorem VEnv.LE.rfl {env : VEnv} : env ≤ env := ⟨id, id⟩ 42 | 43 | theorem VEnv.LE.trans {a b c : VEnv} (h1 : a ≤ b) (h2 : b ≤ c) : a ≤ c := 44 | ⟨h2.1 ∘ h1.1, h2.2 ∘ h1.2⟩ 45 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/Env.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Basic 2 | import Lean4Lean.Theory.VDecl 3 | import Lean4Lean.Theory.Quot 4 | import Lean4Lean.Theory.Inductive 5 | 6 | namespace Lean4Lean 7 | 8 | def VDefVal.WF (env : VEnv) (ci : VDefVal) : Prop := env.HasType ci.uvars [] ci.value ci.type 9 | 10 | inductive VDecl.WF : VEnv → VDecl → VEnv → Prop where 11 | | block : 12 | env.addConst n none = some env' → 13 | VDecl.WF env (.block n) env' 14 | | axiom : 15 | ci.WF env → 16 | env.addConst ci.name (some ci.toVConstant) = some env' → 17 | VDecl.WF env (.axiom ci) env' 18 | | def : 19 | ci.WF env → 20 | env.addConst ci.name (some ci.toVConstant) = some env' → 21 | VDecl.WF env (.def ci) (env'.addDefEq ci.toDefEq) 22 | | opaque : 23 | ci.WF env → 24 | env.addConst ci.name (some ci.toVConstant) = some env' → 25 | VDecl.WF env (.opaque ci) env' 26 | | example : 27 | ci.WF env → 28 | VDecl.WF env (.example ci) env 29 | | quot : 30 | env.QuotReady → 31 | env.addQuot = some env' → 32 | VDecl.WF env .quot env' 33 | | induct : 34 | decl.WF env → 35 | env.addInduct decl = some env' → 36 | VDecl.WF env (.induct decl) env' 37 | 38 | inductive VEnv.WF' : List VDecl → VEnv → Prop where 39 | | empty : VEnv.WF' [] .empty 40 | | decl {env} : VDecl.WF env d env' → env.WF' ds → env'.WF' (d::ds) 41 | 42 | def VEnv.WF (env : VEnv) : Prop := ∃ ds, VEnv.WF' ds env 43 | -------------------------------------------------------------------------------- /Lean4Lean/Std/NodupKeys.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Data.List.Perm 2 | import Batteries.Tactic.SeqFocus 3 | 4 | namespace List 5 | def NodupKeys (l : List (α × β)) : Prop := Nodup (l.map (·.1)) 6 | 7 | theorem nodupKeys_iff_pairwise {l : List (α × β)} : 8 | NodupKeys l ↔ Pairwise (fun a b => a.1 ≠ b.1) l := pairwise_map 9 | 10 | theorem NodupKeys.filter {l : List (α × β)} (p : α × β → Bool) 11 | (H : NodupKeys l) : NodupKeys (l.filter p) := pairwise_map.2 <| (pairwise_map.1 H).filter _ 12 | 13 | @[simp] theorem nodupKeys_cons : NodupKeys (a::l) ↔ (∀ a' ∈ l, a.1 ≠ a'.1) ∧ NodupKeys l := by 14 | simp [nodupKeys_iff_pairwise] 15 | 16 | theorem Perm.nodupKeys_iff {l₁ l₂ : List (α × β)} (h : l₁ ~ l₂) : 17 | NodupKeys l₁ ↔ NodupKeys l₂ := (h.map _).nodup_iff 18 | 19 | theorem Perm.lookup_eq [BEq α] [LawfulBEq α] 20 | {l₁ l₂ : List (α × β)} (h : l₁ ~ l₂) (nd : l₁.NodupKeys) : l₁.lookup a = l₂.lookup a := by 21 | induction h with 22 | | nil => rfl 23 | | cons _ _ ih => simp [lookup]; split <;> [rfl; exact ih (pairwise_cons.1 nd).2] 24 | | swap _ h ih => 25 | simp at nd 26 | simp [lookup]; split <;> split <;> rename_i h1 _ h2 <;> try rfl 27 | simp at h1 h2; cases nd.1.1 (by rw [← h1, h2]) 28 | | trans h1 _ ih1 ih2 => exact (ih1 nd).trans (ih2 (h1.nodupKeys_iff.1 nd)) 29 | 30 | theorem NodupKeys.lookup_eq_some [BEq α] [LawfulBEq α] 31 | {l : List (α × β)} (nd : l.NodupKeys) : l.lookup a = some b ↔ (a, b) ∈ l := by 32 | induction l with 33 | | nil => simp 34 | | cons p l ih => 35 | let (a', b') := p 36 | simp [lookup] at nd ⊢; split <;> rename_i h 37 | · simp at h; subst a'; simp [eq_comm] 38 | exact (nomatch nd.1 _ _ · rfl) 39 | · simp at h; simp [h, ih nd.2] 40 | -------------------------------------------------------------------------------- /Lean4Lean/Std/PersistentHashMap.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.SMap 2 | import Std.Data.HashMap.Lemmas 3 | import Lean4Lean.Verify.Axioms 4 | 5 | namespace Lean.PersistentHashMap 6 | 7 | @[simp] theorem toList'_empty [BEq α] [Hashable α] : 8 | (.empty : PersistentHashMap α β).toList' = [] := by 9 | have this n : @Node.toList' α β (.entries ⟨.replicate n .null⟩) = [] := by 10 | simp [Node.toList'] 11 | induction n <;> simp [*, List.replicate_succ] 12 | apply this 13 | 14 | @[simp] theorem toList'_empty' [BEq α] [Hashable α] : 15 | ({} : PersistentHashMap α β).toList' = [] := toList'_empty 16 | 17 | theorem find?_isSome {α β} [BEq α] [Hashable α] 18 | (m : PersistentHashMap α β) (a : α) : m.contains a = (m.find? a).isSome := findAux_isSome .. 19 | 20 | theorem WF.nodupKeys [BEq α] [Hashable α] 21 | [LawfulBEq α] [LawfulHashable α] 22 | {m : PersistentHashMap α β} (h : WF m) : m.toList'.NodupKeys := by 23 | induction h with 24 | | empty => simp; exact .nil 25 | | insert h1 ih => 26 | refine (h1.toList'_insert ..).nodupKeys_iff.2 (List.nodupKeys_cons.2 ⟨?_, ih.filter _⟩) 27 | rintro _ h3 rfl 28 | simpa using (List.mem_filter.1 h3).2 29 | 30 | variable [BEq α] [LawfulBEq α] [Hashable α] [LawfulHashable α] in 31 | theorem WF.find?_insert {s : PersistentHashMap α β} (h : s.WF) : 32 | (s.insert k v).find? x = if k == x then some v else s.find? x := by 33 | rw [h.insert.find?_eq, h.find?_eq, BEq.comm, 34 | (h.toList'_insert ..).lookup_eq h.insert.nodupKeys, List.lookup] 35 | cases eq : x == k <;> simp 36 | induction s.toList' with 37 | | nil => rfl 38 | | cons kv l ih => 39 | simp [List.filter]; split <;> simp [List.lookup, *] 40 | split <;> [skip; rfl] 41 | rename_i h1 _ h2 42 | simp at h1 h2; simp [h1, h2] at eq 43 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/QuotLemmas.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Env 2 | import Lean4Lean.Theory.Typing.Meta 3 | 4 | namespace Lean4Lean 5 | namespace VEnv 6 | 7 | theorem addQuot_WF (henv : Ordered env) (hq : QuotReady env) : 8 | addQuot env = some env' → Ordered env' := by 9 | refine with_addConst (cis := [.const _ eqConst]) (hcis := ⟨hq, ⟨⟩⟩) henv 10 | (fun _ => ⟨_, by type_tac⟩) fun henv => ?_ 11 | refine with_addConst henv (fun ⟨quot, _⟩ => ⟨_, by type_tac⟩) fun henv => ?_ 12 | refine with_addConst henv (fun ⟨_, quot, eq, _⟩ => ⟨_, by type_tac⟩) fun henv => ?_ 13 | refine with_addConst henv (fun ⟨_, _, quot, _⟩ => ⟨_, by type_tac⟩) fun henv => ?_ 14 | rintro ⟨_, _, _, _, eq, _⟩ ⟨⟩; exact .defeq henv ⟨by type_tac, by type_tac⟩ 15 | 16 | section 17 | variable (henv : addQuot env = some env') include henv 18 | 19 | theorem addQuot_objs : env'.HasObjects [.defeq quotDefEq, .const `Quot.ind quotIndConst, 20 | .const `Quot.lift quotLiftConst, .const `Quot.mk quotMkConst, .const `Quot quotConst] := by 21 | let ⟨env, h, henv⟩ := HasObjects.bind_const (ls := []) trivial henv 22 | let ⟨env, h, henv⟩ := HasObjects.bind_const h henv 23 | let ⟨env, h, henv⟩ := HasObjects.bind_const h henv 24 | obtain ⟨env, h, ⟨⟩⟩ := HasObjects.bind_const h henv 25 | exact HasObjects.defeq (df := quotDefEq) h 26 | 27 | theorem addQuot_quot : env'.constants ``Quot = quotConst := (addQuot_objs henv).2.2.2.2.1 28 | theorem addQuot_quotMk : env'.constants ``Quot.mk = quotMkConst := (addQuot_objs henv).2.2.2.1 29 | theorem addQuot_quotLift : env'.constants ``Quot.lift = quotLiftConst := (addQuot_objs henv).2.2.1 30 | theorem addQuot_quotInd : env'.constants ``Quot.ind = quotIndConst := (addQuot_objs henv).2.1 31 | theorem addQuot_defeq : env'.defeqs quotDefEq := (addQuot_objs henv).1 32 | 33 | end 34 | -------------------------------------------------------------------------------- /Lean4Lean.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Environment 2 | 3 | /- 4 | open Lean Elab Command 5 | 6 | namespace CollectPartial 7 | 8 | structure State where 9 | visited : NameSet := {} 10 | opaques : Array (List Expr) := #[] 11 | 12 | abbrev M := ReaderT Environment $ StateT State MetaM 13 | 14 | partial def collect (stack : List Expr) (c : Name) : M Unit := do 15 | let collectExpr (e : Expr) : M Unit := e.getUsedConstants.forM (collect (mkConst c::stack)) 16 | let s ← get 17 | unless s.visited.contains c do 18 | modify fun s => { s with visited := s.visited.insert c } 19 | let env ← read 20 | match env.find? c with 21 | | some (ConstantInfo.ctorInfo _) 22 | | some (ConstantInfo.recInfo _) 23 | | some (ConstantInfo.inductInfo _) 24 | | some (ConstantInfo.quotInfo _) => 25 | pure () 26 | | some (ConstantInfo.defnInfo v) 27 | | some (ConstantInfo.thmInfo v) => 28 | unless ← Meta.isProp v.type do collectExpr v.value 29 | | some (ConstantInfo.axiomInfo v) 30 | | some (ConstantInfo.opaqueInfo v) => 31 | unless ← Meta.isProp v.type do 32 | modify fun s => { s with opaques := s.opaques.push (mkConst c:: stack) } 33 | | none => 34 | modify fun s => { s with opaques := s.opaques.push (mkConst c:: stack) } 35 | 36 | end CollectPartial 37 | 38 | elab "#print" "partial" name:ident : command => do 39 | let constName ← resolveGlobalConstNoOverloadWithInfo name 40 | let env ← getEnv 41 | let (_, s) ← liftTermElabM <| ((CollectPartial.collect [] constName).run env).run {} 42 | if s.opaques.isEmpty then 43 | logInfo m!"'{constName}' does not use any partial definitions" 44 | else 45 | logInfo m!"'{constName}' depends on opaques: {s.opaques.toList}" 46 | 47 | #print partial Environment.addDecl' 48 | -/ 49 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/Meta.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Lemmas 2 | 3 | namespace Lean4Lean 4 | namespace VEnv 5 | 6 | theorem with_addConst {P : Prop} {env env' : VEnv} {cis : List VObject} 7 | (henv : Ordered env) 8 | (hci : env.HasObjects cis → ci.WF env) 9 | (IH : ∀ {env1}, Ordered env1 → 10 | env1.HasObjects (.const n (some ci) :: cis) → f env1 = some env' → P) 11 | (hcis : env.HasObjects cis) 12 | (H : (env.addConst n (some ci) >>= f) = some env') : P := by 13 | let ⟨env1, h1, henv1⟩ := Option.bind_eq_some_iff.1 H 14 | refine IH (.const henv (by rintro _ ⟨⟩; exact hci hcis) h1) (hcis.const h1) henv1 15 | 16 | theorem Lookup.zero' (eq : A.lift = ty') : 17 | Lookup (A::Γ) 0 ty' := eq ▸ .zero 18 | theorem Lookup.succ' (h : Lookup Γ n ty) (eq : ty.lift = ty') : 19 | Lookup (A::Γ) (n+1) ty' := eq ▸ .succ h 20 | 21 | syntax "lookup_tac" : tactic 22 | macro_rules | `(tactic| lookup_tac) => `(tactic| 23 | first 24 | | refine Lookup.zero' ?_; exact rfl 25 | | refine Lookup.succ' (ty := ?_) ?_ ?_ <;> [skip; lookup_tac; exact rfl] 26 | ) 27 | 28 | theorem HasType.app' (h1 : HasType env U Γ f (.forallE A B)) (h2 : HasType env U Γ a A) 29 | (eq : B.inst a = B') : HasType env U Γ (.app f a) B' := eq ▸ .appDF h1 h2 30 | theorem HasType.const' 31 | (h1 : constants env c = some (some ci)) 32 | (h2 : ∀ (l : VLevel), l ∈ ls → VLevel.WF uvars l) 33 | (h3 : List.length ls = ci.uvars) 34 | (eq : .instL ls ci.type = ty') : 35 | HasType env uvars Γ (.const c ls) ty' := eq ▸ .const h1 h2 h3 36 | 37 | syntax "type_tac" : tactic -- TODO: write an actual tactic 38 | macro_rules | `(tactic| type_tac) => `(tactic| 39 | first 40 | | refine HasType.forallE (u := ?_) (v := ?_) ?_ ?_ <;> [skip; skip; type_tac; type_tac] 41 | | refine HasType.sort ?_; decide 42 | | refine HasType.bvar ?_; lookup_tac 43 | | refine HasType.app' (A := ?_) (B := ?_) ?_ ?_ ?_ <;> [skip; skip; type_tac; type_tac; exact rfl] 44 | | refine HasType.const' (ci := ?_) ?_ ?_ ?_ ?_ <;> [skip; assumption; decide; decide; exact rfl] 45 | | refine HasType.lam (u := ?_) ?_ ?_ <;> [skip; type_tac; type_tac] 46 | ) 47 | -------------------------------------------------------------------------------- /Lean4Lean/EquivManager.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Data.UnionFind.Basic 2 | import Lean4Lean.PtrEq 3 | 4 | namespace Lean4Lean 5 | open Lean 6 | 7 | abbrev EquivManager.NodeRef := Nat 8 | open EquivManager 9 | 10 | structure EquivManager where 11 | uf : Batteries.UnionFind := {} 12 | toNodeMap : ExprMap NodeRef := {} 13 | 14 | namespace EquivManager 15 | 16 | def find (n : NodeRef) : StateM EquivManager NodeRef := fun m => 17 | if h : n < m.uf.size then 18 | let ⟨uf, root, _⟩ := m.uf.find ⟨n, h⟩ 19 | (root, { m with uf }) 20 | else (n, m) 21 | 22 | def merge (m : EquivManager) (n1 n2 : NodeRef) : EquivManager := 23 | if h1 : n1 < m.uf.size then 24 | if h2 : n2 < m.uf.size then 25 | { m with uf := m.uf.union ⟨n1, h1⟩ ⟨n2, h2⟩ } 26 | else m 27 | else m 28 | 29 | def toNode (e : Expr) : StateM EquivManager NodeRef := fun m => do 30 | if let some r := m.toNodeMap[e]? then 31 | return (r, m) 32 | let { uf, toNodeMap } := m 33 | let r := uf.size 34 | (r, { uf := uf.push, toNodeMap := toNodeMap.insert e r }) 35 | 36 | variable (useHash : Bool) in 37 | def isEquiv (e1 e2 : Expr) : StateM EquivManager Bool := do 38 | if ptrEqExpr e1 e2 then return true 39 | if useHash && e1.hash != e2.hash then return false 40 | if e1.isBVar && e2.isBVar then return e1.bvarIdx! == e2.bvarIdx! 41 | let r1 ← find (← toNode e1) 42 | let r2 ← find (← toNode e2) 43 | if r1 == r2 then return true 44 | let result ← 45 | match e1, e2 with 46 | | .const c1 l1, .const c2 l2 => pure <| c1 == c2 && l1 == l2 47 | | .mvar a1, .mvar a2 | .fvar a1, .fvar a2 48 | | .sort a1, .sort a2 | .lit a1, .lit a2 => pure <| a1 == a2 49 | | .app f1 a1, .app f2 a2 => isEquiv f1 f2 <&&> isEquiv a1 a2 50 | | .lam _ d1 b1 _, .lam _ d2 b2 _ => isEquiv d1 d2 <&&> isEquiv b1 b2 51 | | .mdata _ a1, .mdata _ a2 => isEquiv a1 a2 52 | | .forallE _ d1 b1 _, .forallE _ d2 b2 _ => isEquiv d1 d2 <&&> isEquiv b1 b2 53 | | .proj _ i1 e1, .proj _ i2 e2 => pure (i1 == i2) <&&> isEquiv e1 e2 54 | | .letE _ t1 v1 b1 _, .letE _ t2 v2 b2 _ => isEquiv t1 t2 <&&> isEquiv v1 v2 <&&> isEquiv b1 b2 55 | | _, _ => return false 56 | if result then 57 | modify (merge · r1 r2) 58 | return result 59 | termination_by e1 60 | 61 | def addEquiv (m : EquivManager) (e1 e2 : Expr) : EquivManager := 62 | let (r1, m) := toNode e1 m 63 | let (r2, m) := toNode e2 m 64 | merge m r1 r2 65 | -------------------------------------------------------------------------------- /Lean4Lean/Std/ToExpr.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Kyle Miller. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Kyle Miller 5 | -/ 6 | import Lean 7 | 8 | /-! # `ToExpr` instances for Mathlib 9 | 10 | This module should be imported by any module that intends to define `ToExpr` instances. 11 | It provides necessary dependencies (the `Lean.ToLevel` class) and it also overrides the instances 12 | that come from core Lean 4 that do not handle universe polymorphism. 13 | (See the module `Lean.ToExpr` for the instances that are overridden.) 14 | 15 | In addition, we provide some additional `ToExpr` instances for core definitions. 16 | -/ 17 | 18 | set_option autoImplicit true 19 | 20 | namespace Mathlib 21 | open Lean 22 | 23 | deriving instance ToExpr for Int 24 | 25 | deriving instance ToExpr for ULift 26 | 27 | /-- Hand-written instance since `PUnit` is a `Sort` rather than a `Type`. -/ 28 | instance [ToLevel.{u}] : ToExpr PUnit.{u+1} where 29 | toExpr _ := mkConst ``PUnit.unit [toLevel.{u+1}] 30 | toTypeExpr := mkConst ``PUnit [toLevel.{u+1}] 31 | 32 | deriving instance ToExpr for String.Pos 33 | deriving instance ToExpr for Substring 34 | deriving instance ToExpr for SourceInfo 35 | deriving instance ToExpr for Syntax.Preresolved 36 | deriving instance ToExpr for Syntax 37 | 38 | open DataValue in 39 | /-- Core of a hand-written `ToExpr` handler for `MData`. 40 | Uses the `KVMap.set*` functions rather than going into the internals 41 | of the `KVMap` data structure. -/ 42 | private def toExprMData (md : MData) : Expr := Id.run do 43 | let mut e := mkConst ``MData.empty 44 | for (k, v) in md do 45 | let k := toExpr k 46 | e := match v with 47 | | ofString v => mkApp3 (mkConst ``KVMap.setString) e k (mkStrLit v) 48 | | ofBool v => mkApp3 (mkConst ``KVMap.setBool) e k (toExpr v) 49 | | ofName v => mkApp3 (mkConst ``KVMap.setName) e k (toExpr v) 50 | | ofNat v => mkApp3 (mkConst ``KVMap.setNat) e k (mkNatLit v) 51 | | ofInt v => mkApp3 (mkConst ``KVMap.setInt) e k (toExpr v) 52 | | ofSyntax v => mkApp3 (mkConst ``KVMap.setSyntax) e k (toExpr v) 53 | return e 54 | 55 | instance : ToExpr MData where 56 | toExpr := toExprMData 57 | toTypeExpr := mkConst ``MData 58 | 59 | deriving instance ToExpr for FVarId 60 | deriving instance ToExpr for MVarId 61 | deriving instance ToExpr for LevelMVarId 62 | deriving instance ToExpr for Level 63 | deriving instance ToExpr for BinderInfo 64 | deriving instance ToExpr for Literal 65 | deriving instance ToExpr for Expr 66 | 67 | end Mathlib 68 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VEnv 2 | 3 | namespace Lean4Lean 4 | 5 | inductive Lookup : List VExpr → Nat → VExpr → Prop where 6 | | zero : Lookup (ty::Γ) 0 ty.lift 7 | | succ : Lookup Γ n ty → Lookup (A::Γ) (n+1) ty.lift 8 | 9 | namespace VEnv 10 | 11 | section 12 | set_option hygiene false 13 | local notation:65 Γ " ⊢ " e " : " A:30 => IsDefEq Γ e e A 14 | local notation:65 Γ " ⊢ " e1 " ≡ " e2 " : " A:30 => IsDefEq Γ e1 e2 A 15 | variable (env : VEnv) (uvars : Nat) 16 | 17 | inductive IsDefEq : List VExpr → VExpr → VExpr → VExpr → Prop where 18 | | bvar : Lookup Γ i A → Γ ⊢ .bvar i : A 19 | | symm : Γ ⊢ e ≡ e' : A → Γ ⊢ e' ≡ e : A 20 | | trans : Γ ⊢ e₁ ≡ e₂ : A → Γ ⊢ e₂ ≡ e₃ : A → Γ ⊢ e₁ ≡ e₃ : A 21 | | sortDF : 22 | l.WF uvars → l'.WF uvars → l ≈ l' → 23 | Γ ⊢ .sort l ≡ .sort l' : .sort (.succ l) 24 | | constDF : 25 | env.constants c = some (some ci) → 26 | (∀ l ∈ ls, l.WF uvars) → 27 | (∀ l ∈ ls', l.WF uvars) → 28 | ls.length = ci.uvars → 29 | List.Forall₂ (· ≈ ·) ls ls' → 30 | Γ ⊢ .const c ls ≡ .const c ls' : ci.type.instL ls 31 | | appDF : 32 | Γ ⊢ f ≡ f' : .forallE A B → 33 | Γ ⊢ a ≡ a' : A → 34 | Γ ⊢ .app f a ≡ .app f' a' : B.inst a 35 | | lamDF : 36 | Γ ⊢ A ≡ A' : .sort u → 37 | A::Γ ⊢ body ≡ body' : B → 38 | Γ ⊢ .lam A body ≡ .lam A' body' : .forallE A B 39 | | forallEDF : 40 | Γ ⊢ A ≡ A' : .sort u → 41 | A::Γ ⊢ body ≡ body' : .sort v → 42 | Γ ⊢ .forallE A body ≡ .forallE A' body' : .sort (.imax u v) 43 | | defeqDF : Γ ⊢ A ≡ B : .sort u → Γ ⊢ e1 ≡ e2 : A → Γ ⊢ e1 ≡ e2 : B 44 | | beta : 45 | A::Γ ⊢ e : B → Γ ⊢ e' : A → 46 | Γ ⊢ .app (.lam A e) e' ≡ e.inst e' : B.inst e' 47 | | eta : 48 | Γ ⊢ e : .forallE A B → 49 | Γ ⊢ .lam A (.app e.lift (.bvar 0)) ≡ e : .forallE A B 50 | | proofIrrel : 51 | Γ ⊢ p : .sort .zero → Γ ⊢ h : p → Γ ⊢ h' : p → 52 | Γ ⊢ h ≡ h' : p 53 | | extra : 54 | env.defeqs df → (∀ l ∈ ls, l.WF uvars) → ls.length = df.uvars → 55 | Γ ⊢ df.lhs.instL ls ≡ df.rhs.instL ls : df.type.instL ls 56 | 57 | end 58 | 59 | def HasType (env : VEnv) (U : Nat) (Γ : List VExpr) (e A : VExpr) : Prop := 60 | IsDefEq env U Γ e e A 61 | 62 | def IsType (env : VEnv) (U : Nat) (Γ : List VExpr) (A : VExpr) : Prop := 63 | ∃ u, env.HasType U Γ A (.sort u) 64 | 65 | def IsDefEqU (env : VEnv) (U : Nat) (Γ : List VExpr) (e₁ e₂ : VExpr) := 66 | ∃ A, env.IsDefEq U Γ e₁ e₂ A 67 | 68 | end VEnv 69 | 70 | def VExpr.WF (env : VEnv) (U : Nat) (Γ : List VExpr) (e : VExpr) := env.IsDefEqU U Γ e e 71 | 72 | def VConstant.WF (env : VEnv) (ci : VConstant) : Prop := env.IsType ci.uvars [] ci.type 73 | 74 | def VDefEq.WF (env : VEnv) (df : VDefEq) : Prop := 75 | env.HasType df.uvars [] df.lhs df.type ∧ env.HasType df.uvars [] df.rhs df.type 76 | -------------------------------------------------------------------------------- /divergences.md: -------------------------------------------------------------------------------- 1 | # Divergences from the Lean Kernel 2 | 3 | This is a list of places where lean4lean deliberately has different behavior from the kernel. Unless specified here, any divergence between lean4lean and [lean4](https://github.com/leanprover/lean4/tree/master/src/kernel) is a bug. 4 | 5 | * [`Lean4Lean.TypeChecker.Inner.reduceNative`](Lean4Lean/TypeChecker.lean): Lean4lean does not support reduction of `reduceBool`. This would involve implementing verified compilation, which while possible would be an additional chunk of work comparable to this entire repo. 6 | * [`Lean4Lean.Environment.checkPrimitiveDef`](Lean4Lean/Primitive.lean), `checkPrimitiveInductive`: Lean does not check that primitives are declared with the correct types and definitional behavior, except in the case of `Eq` which is used in the declaration of `Quot`. This is required for soundness, but Lean is able to get away with it because Lean ships its prelude and using an alternative prelude is not supported. 7 | * [`Lean4Lean.TypeChecker.Inner.inferType'`](Lean4Lean/TypeChecker.lean), literal case: The original code was not checking that the literal type actually exists. Again, this is okay provided that the prelude is trusted. 8 | * [`Lean.Level.normalize'`](Lean4Lean/Level.lean), `isEquiv'`, `geq'`: Lean4lean implements an experimental new algorithm for level normalization, which is complete for level algebra. We may in the future use a hybrid approach avoid the performance cost in the common case. 9 | * [`Lean4Lean.addDefinition`](Lean4Lean/Environment.lean), `Lean4Lean.addTheorem`: two calls ([1](https://github.com/leanprover/lean4/blob/v4.23.0/src/kernel/environment.cpp#L183) [2](https://github.com/leanprover/lean4/blob/v4.23.0/src/kernel/environment.cpp#L203)) are redundant and have been removed. 10 | * [`Lean4Lean.TypeChecker.Inner.inferLambda`](Lean4Lean/TypeChecker.lean), `inferLet`: lean4lean does the `ensureSort` call before extending the context, while [`infer_lambda`](https://github.com/leanprover/lean4/blob/v4.23.0/src/kernel/type_checker.cpp#L124-L126) does it afterward. It's not clear whether this is actually unsound but it would require some very weird invariants to justify having unchecked things in the local context and hoping that they won't be used in the typing proof of that same expression. 11 | * [`Lean4Lean.TypeChecker.Inner.inferLet`](Lean4Lean/TypeChecker.lean): The implementation is simplified to fix bug [lean4#10475](https://github.com/leanprover/lean4/issues/10475). This will go away once the upstream bugfix lands on stable. 12 | * [`Lean4Lean.checkConstantVal`](Lean4Lean/Environment.lean): The original implementation would call `check` which sets the level params and then unsets them afterward, and then `ensure_sort` would run in a context without any level params. In lean4lean the monad is parameterized over level params, so they remain the same across the two calls. 13 | -------------------------------------------------------------------------------- /Lean4Lean/Std/SMap.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.SMap 2 | import Std.Data.HashMap.Lemmas 3 | import Lean4Lean.Std.HashMap 4 | import Lean4Lean.Std.PersistentHashMap 5 | 6 | namespace Lean.SMap 7 | 8 | variable [BEq α] [Hashable α] 9 | structure WF (s : SMap α β) where 10 | map₂ : s.map₂.WF 11 | stage : s.stage₁ → s.map₂ = .empty 12 | disjoint : s.map₂.contains a → ¬a ∈ s.map₁ 13 | 14 | theorem WF.empty : WF ({} : SMap α β) := ⟨.empty, fun _ => rfl, by simp⟩ 15 | 16 | protected nonrec theorem WF.insert [LawfulBEq α] [LawfulHashable α] {s : SMap α β} 17 | (h : s.WF) (k : α) (v : β) (hn : s.find? k = none) : (s.insert k v).WF := by 18 | unfold insert; split 19 | · refine ⟨h.map₂, h.stage, fun _ h' => ?_⟩ 20 | have := h.stage 21 | simp_all [find?, PersistentHashMap.find?_isSome, PersistentHashMap.WF.empty.find?_eq] 22 | · refine ⟨h.map₂.insert .., nofun, fun h1 => ?_⟩ 23 | simp_all [PersistentHashMap.find?_isSome, h.map₂.find?_insert] 24 | revert ‹_›; split 25 | · simp_all [find?] 26 | · exact h.disjoint ∘ by simp [PersistentHashMap.find?_isSome] 27 | 28 | variable [LawfulBEq α] [LawfulHashable α] in 29 | theorem WF.find?_insert {s : SMap α β} (h : s.WF) : 30 | (s.insert k v).find? x = if k == x then some v else s.find? x := by 31 | unfold insert; split <;> simp [find?] 32 | · exact Std.HashMap.getElem?_insert (α := α) 33 | · rw [h.map₂.find?_insert]; split <;> rfl 34 | 35 | noncomputable def toList' [BEq α] [Hashable α] (m : SMap α β) : 36 | List (α × β) := m.map₂.toList' ++ m.map₁.toList 37 | 38 | open scoped _root_.List in 39 | theorem WF.toList'_insert {α β} [BEq α] [LawfulBEq α] [Hashable α] [LawfulHashable α] 40 | {m : SMap α β} (wf : WF m) (a : α) (b : β) 41 | (h : m.find? a = none) : 42 | (m.insert a b).toList' ~ (a, b) :: m.toList' := by 43 | unfold insert; split <;> simp [toList'] 44 | · have : EquivBEq α := inferInstance; clear ‹LawfulBEq _› 45 | have := wf.stage rfl; simp [find?] at this h; subst this; simp 46 | refine (List.filter_eq_self.2 ?_ ▸ Std.HashMap.insert_toList (α := α) .. :) 47 | rintro ⟨a', b⟩ h 48 | refine Decidable.by_contra fun h2 => ?_ 49 | simp at h2 50 | have := Std.HashMap.getElem?_eq_some_iff_exists_beq_and_mem_toList.2 ⟨_, h2, h⟩ 51 | exact ‹¬_› (Std.HashMap.mem_of_getElem? this) 52 | · refine .append_right (l₂ := _::_) _ ?_ 53 | refine (List.filter_eq_self.2 ?_ ▸ wf.map₂.toList'_insert .. :) 54 | rintro ⟨a', b⟩ h' 55 | have := wf.map₂.find?_eq a 56 | simp_all [find?]; rintro rfl 57 | exact h.1 _ _ h' rfl 58 | 59 | theorem WF.find?_eq {α β} [BEq α] [Hashable α] [LawfulBEq α] [LawfulHashable α] 60 | {m : SMap α β} (wf : WF m) (a : α) : m.find? a = m.toList'.lookup a := by 61 | simp [find?]; split 62 | · cases wf.stage rfl 63 | simp [toList'] 64 | exact Std.HashMap.getElem?_eq_lookup_toList .. 65 | · simp [toList', List.lookup_append, wf.map₂.find?_eq] 66 | cases List.lookup a .. <;> simp [Std.HashMap.getElem?_eq_lookup_toList] 67 | 68 | theorem WF.find?'_eq_find? {α β} [BEq α] [Hashable α] [LawfulBEq α] [LawfulHashable α] 69 | {m : SMap α β} (wf : WF m) (a : α) : m.find?' a = m.find? a := by 70 | simp [find?, find?']; split; · rfl 71 | rename_i m₁ m₂ 72 | cases e1 : m₁[a]? <;> cases e2 : m₂.find? a <;> simp 73 | cases wf.disjoint (by simp [PersistentHashMap.find?_isSome, e2]) (Std.HashMap.mem_of_getElem? e1) 74 | -------------------------------------------------------------------------------- /Lean4Lean/Inductive/Reduce.lean: -------------------------------------------------------------------------------- 1 | import Lean.Structure 2 | import Lean4Lean.Expr 3 | import Lean4Lean.Environment.Basic 4 | 5 | namespace Lean4Lean 6 | open Lean hiding Environment 7 | open Kernel 8 | 9 | section 10 | variable [Monad m] (env : Environment) 11 | (whnf : Expr → m Expr) (inferType : Expr → m Expr) (isDefEq : Expr → Expr → m Bool) 12 | 13 | def getFirstCtor (dName : Name) : Option Name := do 14 | let some (.inductInfo info) := env.find? dName | none 15 | info.ctors.head? 16 | 17 | def mkNullaryCtor (type : Expr) (nparams : Nat) : Option Expr := 18 | type.withApp fun d args => do 19 | let .const dName ls := d | none 20 | let name ← getFirstCtor env dName 21 | return mkAppRange (.const name ls) 0 nparams args 22 | 23 | def toCtorWhenK (rval : RecursorVal) (e : Expr) : m Expr := do 24 | assert! rval.k 25 | let appType ← whnf (← inferType e) 26 | let .const appTypeI _ := appType.getAppFn | return e 27 | if appTypeI != rval.getMajorInduct then return e 28 | if appType.hasExprMVar then 29 | let appTypeArgs := appType.getAppArgs 30 | for h : i in [rval.numParams:appTypeArgs.size] do 31 | if appTypeArgs[i].hasExprMVar then return e 32 | let some newCtorApp := mkNullaryCtor env appType rval.numParams | return e 33 | unless ← isDefEq appType (← inferType newCtorApp) do return e 34 | return newCtorApp 35 | 36 | def expandEtaStruct (eType e : Expr) : Expr := 37 | eType.withApp fun I args => Id.run do 38 | let .const I ls := I | return e 39 | let some ctor := getFirstCtor env I | return e 40 | let some (.ctorInfo info) := env.find? ctor | unreachable! 41 | let mut result := mkAppRange (.const ctor ls) 0 info.numParams args 42 | for i in [:info.numFields] do 43 | result := .app result (.proj I i e) 44 | pure result 45 | 46 | def toCtorWhenStruct (inductName : Name) (e : Expr) : m Expr := do 47 | if !env.isStructureLike inductName || (e.isConstructorApp?' env).isSome then 48 | return e 49 | let eType ← whnf (← inferType e) 50 | if !eType.getAppFn.isConstOf inductName then return e 51 | if (← whnf (← inferType eType)) == .prop then return e 52 | return expandEtaStruct env eType e 53 | 54 | def getRecRuleFor (rval : RecursorVal) (major : Expr) : Option RecursorRule := do 55 | let .const fn _ := major.getAppFn | none 56 | rval.rules.find? (·.ctor == fn) 57 | 58 | def inductiveReduceRec [Monad m] (env : Environment) (e : Expr) 59 | (whnf : Expr → m Expr) (inferType : Expr → m Expr) (isDefEq : Expr → Expr → m Bool) : 60 | m (Option Expr) := do 61 | let .const recFn ls := e.getAppFn | return none 62 | let some (.recInfo info) := env.find? recFn | return none 63 | let recArgs := e.getAppArgs 64 | let majorIdx := info.getMajorIdx 65 | let some major := recArgs[majorIdx]? | return none 66 | let mut major := major 67 | if info.k then 68 | major ← toCtorWhenK env whnf inferType isDefEq info major 69 | match ← whnf major with 70 | | .lit l => major := l.toConstructor 71 | | e => major ← toCtorWhenStruct env whnf inferType info.getMajorInduct e 72 | let some rule := getRecRuleFor info major | return none 73 | let majorArgs := major.getAppArgs 74 | if rule.nfields > majorArgs.size then return none 75 | if ls.length != info.levelParams.length then return none 76 | let mut rhs := rule.rhs.instantiateLevelParams info.levelParams ls 77 | rhs := mkAppRange rhs 0 info.getFirstIndexIdx recArgs 78 | rhs := mkAppRange rhs (majorArgs.size - rule.nfields) majorArgs.size majorArgs 79 | if majorIdx + 1 < recArgs.size then 80 | rhs := mkAppRange rhs (majorIdx + 1) recArgs.size recArgs 81 | return rhs 82 | 83 | end 84 | -------------------------------------------------------------------------------- /Lean4Lean/Expr.lean: -------------------------------------------------------------------------------- 1 | import Lean.Environment 2 | 3 | namespace Lean 4 | namespace Expr 5 | 6 | def prop : Expr := .sort .zero 7 | 8 | def arrow (d b : Expr) : Expr := .forallE `a d b .default 9 | 10 | def lam0 (ty e : Expr) : Expr := .lam `_ ty e default 11 | 12 | namespace ReplaceImpl 13 | 14 | unsafe abbrev ReplaceT := StateT (PtrMap Expr Expr) 15 | 16 | @[inline] 17 | unsafe def cacheT [Monad m] (key : Expr) (result : Expr) : ReplaceT m Expr := do 18 | modify (·.insert key result) 19 | pure result 20 | 21 | @[specialize] 22 | unsafe def replaceUnsafeT [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : ReplaceT m Expr := do 23 | let rec @[specialize] visit (e : Expr) := do 24 | if let some result := (← get).find? e then 25 | return result 26 | match ← f? e with 27 | | some eNew => cacheT e eNew 28 | | none => match e with 29 | | Expr.forallE _ d b _ => cacheT e <| e.updateForallE! (← visit d) (← visit b) 30 | | Expr.lam _ d b _ => cacheT e <| e.updateLambdaE! (← visit d) (← visit b) 31 | | Expr.mdata _ b => cacheT e <| e.updateMData! (← visit b) 32 | | Expr.letE _ t v b nd => cacheT e <| e.updateLet! (← visit t) (← visit v) (← visit b) nd 33 | | Expr.app f a => cacheT e <| e.updateApp! (← visit f) (← visit a) 34 | | Expr.proj _ _ b => cacheT e <| e.updateProj! (← visit b) 35 | | e => pure e 36 | visit e 37 | 38 | @[inline] 39 | unsafe def replaceUnsafe' [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := 40 | (replaceUnsafeT f? e).run' mkPtrMap 41 | 42 | end ReplaceImpl 43 | 44 | /- TODO: use withPtrAddr, withPtrEq to avoid unsafe tricks above. 45 | We also need an invariant at `State` and proofs for the `uget` operations. -/ 46 | 47 | @[specialize] 48 | def replaceNoCacheT [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := do 49 | match ← f? e with 50 | | some eNew => pure eNew 51 | | none => match e with 52 | | .forallE _ d b _ => 53 | return e.updateForallE! (← replaceNoCacheT f? d) (← replaceNoCacheT f? b) 54 | | .lam _ d b _ => 55 | return e.updateLambdaE! (← replaceNoCacheT f? d) (← replaceNoCacheT f? b) 56 | | .mdata _ b => 57 | return e.updateMData! (← replaceNoCacheT f? b) 58 | | .letE _ t v b nd => 59 | return e.updateLet! 60 | (← replaceNoCacheT f? t) (← replaceNoCacheT f? v) (← replaceNoCacheT f? b) nd 61 | | .app f a => 62 | return e.updateApp! (← replaceNoCacheT f? f) (← replaceNoCacheT f? a) 63 | | .proj _ _ b => 64 | return e.updateProj! (← replaceNoCacheT f? b) 65 | | e => return e 66 | 67 | @[implemented_by ReplaceImpl.replaceUnsafe'] 68 | partial def replaceM [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := 69 | e.replaceNoCacheT f? 70 | 71 | def natZero : Expr := .const ``Nat.zero [] 72 | def natSucc : Expr := .const ``Nat.succ [] 73 | 74 | def isConstructorApp?' (env : Kernel.Environment) (e : Expr) : Option Name := do 75 | let .const fn _ := e.getAppFn | none 76 | let .ctorInfo _ ← env.find? fn | none 77 | return fn 78 | 79 | def natLitToConstructor : Nat → Expr 80 | | 0 => natZero 81 | | n+1 => .app natSucc (.lit (.natVal n)) 82 | 83 | def strLitToConstructor (s : String) : Expr := 84 | let char := .const ``Char [] 85 | let listNil := .app (.const ``List.nil [.zero]) char 86 | let listCons := .app (.const ``List.cons [.zero]) char 87 | let stringMk := .const ``String.mk [] 88 | let charOfNat := .const ``Char.ofNat [] 89 | .app stringMk <| s.foldr (init := listNil) fun c e => 90 | .app (.app listCons <| .app charOfNat (.lit (.natVal c.toNat))) e 91 | 92 | end Expr 93 | 94 | namespace Literal 95 | 96 | def toConstructor : Literal → Expr 97 | | .natVal n => .natLitToConstructor n 98 | | .strVal s => .strLitToConstructor s 99 | 100 | /-- Return the type of a literal value. -/ 101 | def typeName : Literal → Name 102 | | .natVal _ => ``Nat 103 | | .strVal _ => ``String 104 | 105 | end Literal 106 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/VLCtx.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.Expr 2 | import Lean4Lean.Theory.VExpr 3 | 4 | namespace Lean4Lean 5 | 6 | open Lean (FVarId Expr) 7 | 8 | inductive VLocalDecl where 9 | | vlam (type : VExpr) 10 | | vlet (type value : VExpr) 11 | 12 | def VLocalDecl.depth : VLocalDecl → Nat 13 | | .vlam .. => 1 14 | | .vlet .. => 0 15 | 16 | def VLocalDecl.value : VLocalDecl → VExpr 17 | | .vlam .. => .bvar 0 18 | | .vlet _ e => e 19 | 20 | def VLocalDecl.type' : VLocalDecl → VExpr 21 | | .vlam A 22 | | .vlet A _ => A 23 | 24 | def VLocalDecl.type : VLocalDecl → VExpr 25 | | .vlam A => A.lift 26 | | .vlet A _ => A 27 | 28 | def VLocalDecl.lift' : VLocalDecl → Lift → VLocalDecl 29 | | .vlam A, n => .vlam (A.lift' n) 30 | | .vlet A e, n => .vlet (A.lift' n) (e.lift' n) 31 | 32 | def VLocalDecl.liftN : VLocalDecl → Nat → Nat → VLocalDecl 33 | | .vlam A, n, k => .vlam (A.liftN n k) 34 | | .vlet A e, n, k => .vlet (A.liftN n k) (e.liftN n k) 35 | 36 | def VLocalDecl.inst : VLocalDecl → VExpr → (k : Nat := 0) → VLocalDecl 37 | | .vlam A, e₀, k => .vlam (A.inst e₀ k) 38 | | .vlet A e, e₀, k => .vlet (A.inst e₀ k) (e.inst e₀ k) 39 | 40 | def VLocalDecl.instL : VLocalDecl → List VLevel → VLocalDecl 41 | | .vlam A, ls => .vlam (A.instL ls) 42 | | .vlet A e, ls => .vlet (A.instL ls) (e.instL ls) 43 | 44 | def VLCtx := List (Option (FVarId × List FVarId) × VLocalDecl) 45 | 46 | namespace VLCtx 47 | 48 | def bvars : VLCtx → Nat 49 | | [] => 0 50 | | (none, _) :: Δ => bvars Δ + 1 51 | | (some _, _) :: Δ => bvars Δ 52 | 53 | abbrev NoBV (Δ : VLCtx) : Prop := Δ.bvars = 0 54 | 55 | def next : Option (FVarId × List FVarId) → Nat ⊕ FVarId → Option (Nat ⊕ FVarId) 56 | | none, .inl 0 => none 57 | | none, .inl (n+1) => some (.inl n) 58 | | some _, .inl n => some (.inl n) 59 | | none, .inr fv' => some (.inr fv') 60 | | some (fv, _), .inr fv' => if fv == fv' then none else some (.inr fv') 61 | 62 | def find? : VLCtx → Nat ⊕ FVarId → Option (VExpr × VExpr) 63 | | [], _ => none 64 | | (ofv, d) :: Δ, v => 65 | match next ofv v with 66 | | none => some (d.value, d.type) 67 | | some v => do let (e, A) ← find? Δ v; some (e.liftN d.depth, A.liftN d.depth) 68 | 69 | def liftVar (n k : Nat) : Nat ⊕ FVarId → Nat ⊕ FVarId 70 | | .inl i => .inl (if i < k then i else i + n) 71 | | .inr fv => .inr fv 72 | 73 | def varToExpr : Nat ⊕ FVarId → Expr 74 | | .inl i => .bvar i 75 | | .inr fv => .fvar fv 76 | 77 | def vlamName : VLCtx → Nat → Option (Option (FVarId × List FVarId)) 78 | | [], _ => none 79 | | (_, .vlet ..) :: Δ, i 80 | | (_, .vlam ..) :: Δ, i+1 => vlamName Δ i 81 | | (ofv, .vlam ..) :: _, 0 => some ofv 82 | 83 | def fvars (Δ : VLCtx) : List FVarId := Δ.filterMap (·.1.map (·.1)) 84 | 85 | @[simp] theorem fvars_nil : fvars [] = [] := rfl 86 | @[simp] theorem fvars_cons_none {Δ : VLCtx} : fvars ((none, d) :: Δ) = fvars Δ := rfl 87 | @[simp] theorem fvars_cons_some {Δ : VLCtx} : 88 | fvars ((some fv, d) :: Δ) = fv.1 :: fvars Δ := rfl 89 | 90 | def toCtx : VLCtx → List VExpr 91 | | [] => [] 92 | | (_, .vlam ty) :: Δ => ty :: VLCtx.toCtx Δ 93 | | (_, .vlet _ _) :: Δ => VLCtx.toCtx Δ 94 | 95 | def instL (Δ : VLCtx) (ls : List VLevel) : VLCtx := 96 | match Δ with 97 | | [] => [] 98 | | (ofv, d) :: Δ => (ofv, d.instL ls) :: instL Δ ls 99 | 100 | theorem find?_eq_some : (∃ x, Δ.find? (.inr fv) = some x) ↔ fv ∈ fvars Δ := by 101 | induction Δ with simp [find?] | cons d Δ ih 102 | match d with 103 | | (none, _) => simp [next, ← ih]; grind 104 | | (some (fv', _), _) => 105 | simp [next]; rw [@eq_comm _ fv] 106 | by_cases h : fv' == fv <;> simp [h] <;> simp at h <;> simp [h]; grind 107 | 108 | theorem vlamName_mem_fvars : 109 | ∀ {Δ : VLCtx} {i}, Δ.vlamName i = some (some fv) → fv.1 ∈ fvars Δ 110 | | (none, .vlet ..) :: Δ, _, h 111 | | (none, .vlam ..) :: Δ, _+1, h => vlamName_mem_fvars (Δ := Δ) h 112 | | (some _, .vlet ..) :: Δ, _, h 113 | | (some _, .vlam ..) :: Δ, _+1, h => .tail _ <| vlamName_mem_fvars (Δ := Δ) h 114 | | (some _fv, .vlam ..) :: _, 0, rfl => .head _ 115 | 116 | end VLCtx 117 | -------------------------------------------------------------------------------- /Lean4Lean/Std/HashMap.lean: -------------------------------------------------------------------------------- 1 | import Std.Data.HashMap.Lemmas 2 | import Lean 3 | 4 | namespace Std.Internal.List 5 | 6 | namespace Const 7 | 8 | def containsKey [BEq α] (a : α) : List (α × β) → Bool 9 | | [] => false 10 | | ⟨k, _⟩ :: l => k == a || containsKey a l 11 | 12 | def replaceEntry [BEq α] (k : α) (v : β) : List (α × β) → List (α × β) 13 | | [] => [] 14 | | ⟨k', v'⟩ :: l => bif k' == k then ⟨k, v⟩ :: l else ⟨k', v'⟩ :: replaceEntry k v l 15 | 16 | def insertEntry [BEq α] (k : α) (v : β) (l : List (α × β)) : List (α × β) := 17 | bif containsKey k l then replaceEntry k v l else (k, v) :: l 18 | 19 | theorem insertEntry_perm_filter [BEq α] [PartialEquivBEq α] 20 | (k : α) (v : β) {l : List (α × β)} (hl : l.Pairwise (fun a b => (a.1 == b.1) = false)) : 21 | (insertEntry k v l).Perm <| (k, v) :: l.filter (¬k == ·.1) := by 22 | unfold insertEntry 23 | cases eq : containsKey k l <;> simp 24 | · rw [List.filter_eq_self.2] 25 | induction l <;> simp_all [containsKey, BEq.comm]; assumption 26 | · induction l <;> simp_all [-List.pairwise_cons, containsKey, replaceEntry, List.filter, BEq.comm] 27 | revert eq; cases h : k == _ <;> intro eq <;> simp_all 28 | rename_i ih 29 | · exact .trans (.cons _ ih) (.swap ..) 30 | · rw [List.filter_eq_self.2] 31 | intro ⟨a, b⟩ h' 32 | have := hl.1 a b h' 33 | refine Decidable.by_contra fun h2 => ?_ 34 | simp at h2 35 | cases h2.symm.trans (BEq.congr_left h ▸ this) 36 | 37 | end Const 38 | 39 | theorem Const.containsKey_map [BEq α] (l : List ((_ : α) × β)) : 40 | Const.containsKey k (l.map (fun x => (x.1, x.2))) = List.containsKey k l := by 41 | induction l <;> simp [List.containsKey, Const.containsKey, *] 42 | 43 | theorem Const.replaceEntry_map [BEq α] (l : List ((_ : α) × β)) : 44 | Const.replaceEntry k v (l.map (fun x => (x.1, x.2))) = 45 | (List.replaceEntry k v l).map (fun x => (x.1, x.2)) := by 46 | induction l <;> simp [List.replaceEntry, Const.replaceEntry, *] 47 | cases _ == k <;> simp 48 | 49 | theorem Const.insertEntry_map [BEq α] (l : List ((_ : α) × β)) : 50 | Const.insertEntry k v (l.map (fun x => (x.1, x.2))) = 51 | (List.insertEntry k v l).map (fun x => (x.1, x.2)) := by 52 | simp [List.insertEntry, Const.insertEntry, Const.containsKey_map, Const.replaceEntry_map] 53 | cases List.containsKey k l <;> simp 54 | 55 | end Std.Internal.List 56 | 57 | namespace Std.DHashMap 58 | open Std.Internal.List Internal.Raw Internal.Raw₀ 59 | 60 | variable [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] 61 | 62 | theorem insert_perm_insertEntry (m : DHashMap α β) (a b) : 63 | (insert m a b).toList.Perm (insertEntry a b m.toList) := by 64 | simp [toList, toList_eq_toListModel] 65 | exact toListModel_insert (WF.out m.2) 66 | 67 | theorem Const.insert_perm_insertEntry (m : DHashMap α fun _ => β) (a b) : 68 | (Const.toList (insert m a b)).Perm (Const.insertEntry a b (Const.toList m)) := by 69 | simp [Const.toList, Const.toList_eq_toListModel_map, Const.insertEntry_map] 70 | exact (toListModel_insert (WF.out m.2)).map _ 71 | 72 | end Std.DHashMap 73 | 74 | namespace Std.HashMap 75 | open Internal.List 76 | 77 | variable [BEq α] [Hashable α] [LawfulHashable α] 78 | 79 | theorem insert_toList [EquivBEq α] (m : HashMap α β) : 80 | (insert m a b).toList.Perm ((a, b) :: m.toList.filter (¬a == ·.1)) := by 81 | refine (DHashMap.Const.insert_perm_insertEntry m.1 a b).trans ?_ 82 | exact Const.insertEntry_perm_filter _ _ distinct_keys_toList 83 | 84 | theorem getElem?_eq_lookup_toList [LawfulBEq α] (m : HashMap α β) (a : α) : 85 | m[a]? = m.toList.lookup a := by 86 | apply Option.ext fun b => ?_ 87 | rw [← mem_toList_iff_getElem?_eq_some] 88 | have := distinct_keys_toList (m := m); revert this 89 | induction m.toList <;> intro H <;> simp_all [List.lookup] 90 | split <;> simp_all [Prod.ext_iff] 91 | simp [eq_comm] 92 | rw [List.lookup_eq_none_iff.2]; · simp 93 | simp; exact H.1 94 | 95 | theorem mem_of_getElem? [EquivBEq α] {m : HashMap α β} (h : m[a]? = some b) : a ∈ m := by 96 | refine Decidable.by_contra fun h' => ?_ 97 | cases (getElem?_eq_none h').symm.trans h 98 | 99 | end Std.HashMap 100 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/TypeChecker.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.TypeChecker.InferType 2 | import Lean4Lean.Verify.TypeChecker.WHNF 3 | import Lean4Lean.Verify.TypeChecker.IsDefEq 4 | 5 | namespace Lean4Lean 6 | 7 | open Lean hiding Environment Exception 8 | open Kernel 9 | 10 | structure VEnvs where 11 | venv : DefinitionSafety → VEnv 12 | 13 | structure VEnvs.WF (env : Environment) (ves : VEnvs) where 14 | tr : TrEnv safety env (ves.venv safety) 15 | hasPrimitives : VEnv.HasPrimitives (ves.venv safety) 16 | safePrimitives : env.find? n = some ci → 17 | Environment.primitives.contains n → isAsSafeAs .safe ci ∧ ci.levelParams = [] 18 | 19 | namespace TypeChecker 20 | open Inner 21 | 22 | theorem Methods.withFuel.WF : ∀ {n}, (withFuel n).WF 23 | | 0 => 24 | { isDefEqCore _ _ := .throw 25 | whnfCore _ := .throw 26 | whnf _ := .throw 27 | inferType _ _ := .throw } 28 | | n + 1 => 29 | have := withFuel.WF (n := n) 30 | { isDefEqCore h1 h2 := isDefEqCore'.WF h1 h2 _ this 31 | whnfCore h1 := whnfCore'.WF h1 _ this 32 | whnf h1 := whnf'.WF h1 _ this 33 | inferType h1 h2 := inferType'.WF h1 h2 _ this } 34 | 35 | theorem RecM.WF.run {x : RecM α} (H : x.WF c s Q) : (RecM.run x).WF c s Q := 36 | H _ Methods.withFuel.WF 37 | 38 | def VContext.mk' {env : Environment} {ves : VEnvs} (wf : ves.WF env) 39 | (safety : DefinitionSafety := .safe) (lparams : List Name := []) : VContext where 40 | env; safety; lparams 41 | venv := ves.venv safety 42 | hasPrimitives := wf.hasPrimitives 43 | safePrimitives := wf.safePrimitives 44 | trenv := wf.tr 45 | mlctx := .nil 46 | mlctx_wf := trivial 47 | lctx_eq := rfl 48 | 49 | theorem VState.WF.empty {env : Environment} {ves : VEnvs} {wf : ves.WF env} 50 | {safety : DefinitionSafety} {lparams : List Name} : 51 | VState.WF (.mk' wf safety lparams) {} where 52 | trctx := .nil 53 | ngen_wf := nofun 54 | ectx := ⟨[], .refl, trivial, .refl, .empty, nofun⟩ 55 | inferTypeI_wf := .empty 56 | inferTypeC_wf := .empty 57 | whnfCore_wf := .empty 58 | whnf_wf := .empty 59 | 60 | theorem M.WF.run {env : Environment} {ves : VEnvs} (wf : ves.WF env) 61 | {x : M α} {Q} (H : x.WF (.mk' wf safety lparams) {} fun a _ => Q a) : 62 | (M.run env safety {} lparams x).WF Q := by 63 | intro a eq 64 | simp [M.run, Functor.map, Except.map] at eq 65 | split at eq <;> cases eq; rename_i eq 66 | let ⟨_, _, _, _, H⟩ := H .empty _ _ eq 67 | exact H 68 | 69 | nonrec theorem whnf.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 70 | M.WF c s (whnf e) fun e₁ _ => c.TrExpr e₁ e' := 71 | (whnf.WF he).run.mono fun _ _ _ h => h.2 72 | 73 | nonrec theorem inferType.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 74 | M.WF c s (inferType e) fun ty _ => ∃ ty', c.TrTyping e ty e' ty' := 75 | (inferType.WF he).run 76 | 77 | nonrec theorem checkType.WF {c : VContext} {s : VState} (h1 : e.FVarsIn (· ∈ c.vlctx.fvars)) : 78 | M.WF c s (checkType e) fun ty _ => ∃ e' ty', c.TrTyping e ty e' ty' := 79 | (checkType.WF h1).run 80 | 81 | nonrec theorem isDefEq.WF {c : VContext} {s : VState} 82 | (he₁ : c.TrExprS e₁ e₁') (he₂ : c.TrExprS e₂ e₂') : 83 | M.WF c s (isDefEq e₁ e₂) fun b _ => b → c.IsDefEqU e₁' e₂' := 84 | (isDefEq.WF he₁ he₂).run 85 | 86 | nonrec theorem isProp.WF {c : VContext} {s : VState} 87 | (he : c.TrExprS e e') : (isProp e).WF c s fun b _ => b → c.HasType e' (.sort .zero) := 88 | (isProp.WF he).run 89 | 90 | nonrec theorem ensureSort.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 91 | M.WF c s (ensureSort e e₀) fun e1 _ => c.TrExpr e1 e' ∧ ∃ u, e1 = .sort u := 92 | (ensureSortCore.WF he).run.mono fun _ _ _ h => ⟨h.2.1, h.1⟩ 93 | 94 | nonrec theorem ensureForall.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 95 | M.WF c s (ensureForall e) fun e1 _ => 96 | c.TrExpr e1 e' ∧ ∃ name ty body bi, e1 = .forallE name ty body bi := 97 | (ensureForallCore.WF he).run.mono fun _ _ _ h => h.2 98 | 99 | nonrec theorem ensureType.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 100 | M.WF c s (ensureType e) fun e1 _ => ∃ e', c.TrExprS e e' ∧ ∃ u u', e1 = .sort u ∧ 101 | VLevel.ofLevel c.lparams u = some u' ∧ c.HasType e' (.sort u') := by 102 | refine (inferType.WF he).bind fun _ _ _ ⟨_, _, a1, a2, a3⟩ => ?_ 103 | refine (ensureSort.WF a2).mono fun _ _ _ ⟨⟨_, b1, b2⟩, b3⟩ => ?_ 104 | obtain ⟨_, rfl⟩ := b3; let .sort b1 := b1 105 | exact ⟨_, a1, _, _, rfl, b1, a3.defeqU_r c.Ewf c.Δwf b2.symm⟩ 106 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Meta.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Lean.Expr 2 | import Lean4Lean.Std.Control 3 | import Lean4Lean.Std.ToExpr 4 | import Lean4Lean.Theory.VEnv 5 | import Lean4Lean.Inductive.Reduce 6 | 7 | namespace Lean4Lean 8 | 9 | namespace Meta 10 | open Lean Meta Elab Term 11 | 12 | def expandProj (structName : Name) (idx : Nat) (e : Expr) : MetaM Expr := do 13 | let failed {α} : Unit → MetaM α := fun _ => do 14 | throwError "invalid projection{indentExpr (mkProj structName idx e)}" 15 | let ival ← getConstInfoInduct structName 16 | let [ctor] := ival.ctors | failed () 17 | let ctorInfo ← getConstInfoCtor ctor 18 | unless idx < ctorInfo.numFields do failed () 19 | let args : Array Term ← (Array.range ctorInfo.numFields).mapM fun _ => do 20 | `($(mkIdent (← mkFreshId)):ident) 21 | let args' ← args.mapM fun x => `(Lean.Parser.Term.funBinder| $x) 22 | let casesOn := mkCIdent (mkCasesOnName structName) 23 | TermElabM.run' do 24 | let stx ← `($casesOn:ident $(← e.toSyntax) fun $args'* => $(args[idx]!)) 25 | elabTerm stx (← inferType (.proj structName idx e)) 26 | 27 | partial def expandExpr (e : Expr) : MetaM Expr := 28 | Meta.transform e 29 | (pre := fun 30 | | .mdata _ e => return .visit e 31 | | .proj s i e => return .continue (← expandProj s i e) 32 | | e@(.mvar _) => return .continue (← instantiateMVars e) 33 | | _ => return .continue) 34 | 35 | variable (ls : List Name) (fvarToIdx : FVarIdMap Nat) in 36 | partial def ofExpr : Expr → (k :_:= 0) → MetaM VExpr 37 | | .bvar i, _ => return .bvar i 38 | | .sort u, _ => return .sort (← VLevel.ofLevel ls u) 39 | | .const c us, _ => return .const c (← liftM <| us.mapM (VLevel.ofLevel ls)) 40 | | .app fn arg, k => return .app (← ofExpr fn k) (← ofExpr arg k) 41 | | .lam _ ty body _, k => return .lam (← ofExpr ty k) (← ofExpr body (k+1)) 42 | | .forallE _ ty body _, k => return .forallE (← ofExpr ty k) (← ofExpr body (k+1)) 43 | | .mdata _ e, k => ofExpr e k 44 | | .lit l, k => ofExpr l.toConstructor k 45 | | .letE _ _ value body _, k => return (← ofExpr body (k+1)).inst (← ofExpr value k) 46 | | e@(.proj ..), _ => throwError "invalid expression {e}" 47 | | e@(.mvar ..), _ => throwError "expression contains metavariables {e}" 48 | | .fvar e, k => do 49 | if let some i := fvarToIdx.get? e then return .bvar (i+k) 50 | let lctx ← getLCtx 51 | let some e := (do (← lctx.find? e).value?) | throwError "undeclared free var {Expr.fvar e}" 52 | ofExpr e 53 | 54 | deriving instance ToExpr for VLevel 55 | deriving instance ToExpr for VExpr 56 | deriving instance ToExpr for VConstant 57 | deriving instance ToExpr for VDefEq 58 | 59 | def toVExprWrapper (bis : Option (TSyntaxArray ``Parser.Term.funBinder)) 60 | (x : FVarIdMap Nat → TermElabM α) : TermElabM α := do 61 | withLevelNames [] <| 62 | elabFunBinders (bis.getD #[]) none fun xs _ => 63 | withAutoBoundImplicit <| 64 | x <| 65 | xs.foldr (fun fvar (n, m) => (n+1, m.insert fvar.fvarId! n)) (0, ({}:FVarIdMap Nat)) |>.2 66 | 67 | def elabForVExpr (e : Expr) : TermElabM Expr := do 68 | let e ← levelMVarToParam e 69 | if ← logUnassignedUsingErrorInfos (← getMVars e) then throwAbortCommand 70 | expandExpr e 71 | 72 | def toVExprCore (bis : Option (TSyntaxArray ``Parser.Term.funBinder)) 73 | (e : Term) : TermElabM (Nat × VExpr) := do 74 | toVExprWrapper bis fun map => do 75 | let e ← elabForVExpr (← elabTerm e none) 76 | let ls := (← getLevelNames).reverse 77 | return (ls.length, ← ofExpr ls map e) 78 | 79 | syntax "vexpr(" atomic(Parser.Term.funBinder* " ⊢ ")? term ")" : term 80 | 81 | elab_rules : term 82 | | `(vexpr($[$bis* ⊢]? $e:term)) => return toExpr (← toVExprCore bis e).2 83 | 84 | syntax "vconst(" atomic(Parser.Term.funBinder* " ⊢ ")? term ")" : term 85 | 86 | elab_rules : term 87 | | `(vconst($[$bis* ⊢]? $e:term)) => do 88 | let (n, ve) ← toVExprCore bis e 89 | return toExpr (⟨n, ve⟩ : VConstant) 90 | 91 | syntax "vdefeq(" atomic(Parser.Term.funBinder* " ⊢ ")? 92 | atomic(Parser.Term.funBinder* " => ")? term " ≡ " term ")" : term 93 | 94 | elab_rules : term 95 | | `(vdefeq($[$bis* ⊢]? $[$args* =>]? $e₁:term ≡ $e₂:term)) => do 96 | toVExprWrapper bis fun map => do 97 | let (e₁, e₂) ← match args with 98 | | some args => pure (← `(fun $args* => $e₁), ← `(fun $args* => $e₂)) 99 | | none => pure (e₁, e₂) 100 | let e₁ ← elabTerm e₁ none 101 | let ty ← inferType e₁ 102 | let e₂ ← elabTerm e₂ ty 103 | let e₁ ← elabForVExpr e₁ 104 | let e₂ ← elabForVExpr e₂ 105 | let ty ← elabForVExpr ty 106 | let ls := (← getLevelNames).reverse 107 | return toExpr { 108 | uvars := ls.length 109 | lhs := ← ofExpr ls map e₁ 110 | rhs := ← ofExpr ls map e₂ 111 | type := ← ofExpr ls map ty 112 | : VDefEq } 113 | -------------------------------------------------------------------------------- /Lean4Lean/Quot.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Tactic.OpenPrivate 2 | import Lean4Lean.Environment.Basic 3 | import Lean4Lean.Expr 4 | import Lean4Lean.Instantiate 5 | import Lean4Lean.LocalContext 6 | 7 | namespace Lean4Lean 8 | open Lean hiding Environment Exception 9 | open Kernel 10 | 11 | open private Lean.Kernel.Environment.add markQuotInit from Lean.Environment 12 | 13 | abbrev ExprBuildT (m) := ReaderT LocalContext <| ReaderT NameGenerator m 14 | 15 | def ExprBuildT.run [Monad m] (x : ExprBuildT m α) : m α := x {} {} 16 | 17 | instance : MonadLocalNameGenerator (ExprBuildT m) where 18 | withFreshId x c ngen := x ngen.curr c ngen.next 19 | 20 | def checkEqType (env : Environment) : Except Exception Unit := do 21 | let fail {α} (s : String) : Except Exception α := 22 | throw <| .other s!"failed to initialize quot module, {s}" 23 | let .inductInfo info ← env.get ``Eq | fail "environment does not have 'Eq' type" 24 | let [u] := info.levelParams | fail "unexpected number of universe params at 'Eq' type" 25 | let [eqRefl] := info.ctors | fail "unexpected number of constructors for 'Eq' type" 26 | ExprBuildT.run do 27 | withLocalDecl `α (.sort (.param u)) .implicit fun α => do 28 | if info.type != ((← read).mkForall #[α] <| .arrow α <| .arrow α .prop) then 29 | fail "'Eq' has an expected type" 30 | let info ← env.get eqRefl 31 | let [u] := info.levelParams 32 | | fail "unexpected number of universe params at 'Eq' type constructor" 33 | withLocalDecl `α (.sort (.param u)) .implicit fun α => do 34 | withLocalDecl `a α .default fun a => do 35 | if info.type != ((← read).mkForall #[α, a] <| mkApp3 (.const ``Eq [.param u]) α a a) then 36 | fail "unexpected type for 'Eq' type constructor" 37 | 38 | def Environment.addQuot (env : Environment) : Except Exception Environment := do 39 | if env.quotInit then return env 40 | checkEqType env 41 | ExprBuildT.run do 42 | let u := .param `u 43 | withLocalDecl `α (.sort u) .implicit fun α => do 44 | let env ← withLocalDecl `r (.arrow α (.arrow α .prop)) .default fun r => do 45 | -- constant Quot.{u} {α : Sort u} (r : α → α → Prop) : Sort u 46 | let env := env.add <| .quotInfo { 47 | name := ``Quot, kind := .type, levelParams := [`u] 48 | type := (← read).mkForall #[α, r] <| .sort u 49 | } 50 | withLocalDecl `a α .default fun a => do 51 | -- constant Quot.mk.{u} {α : Sort u} (r : α → α → Prop) (a : α) : @Quot.{u} α r 52 | return env.add <| .quotInfo { 53 | name := ``Quot.mk, kind := .ctor, levelParams := [`u] 54 | type := (← read).mkForall #[α, r, a] <| mkApp2 (.const ``Quot [u]) α r 55 | } 56 | withLocalDecl `r (.arrow α (.arrow α .prop)) .implicit fun r => do 57 | let quot_r := mkApp2 (.const ``Quot [u]) α r 58 | withLocalDecl `a α .default fun a => do 59 | let v := .param `v 60 | let env ← withLocalDecl `β (.sort v) .implicit fun β => do 61 | withLocalDecl `f (.arrow α β) .default fun f => do 62 | withLocalDecl `b α .default fun b => do 63 | let rab := mkApp2 r a b 64 | let fa_eq_fb := mkApp3 (.const ``Eq [v]) β (.app f a) (.app f b) 65 | let sanity := (← read).mkForall #[a, b] <| .arrow rab fa_eq_fb 66 | -- constant Quot.lift.{u, v} {α : Sort u} {r : α → α → Prop} {β : Sort v} (f : α → β) : 67 | -- (∀ a b : α, r a b → f a = f b) → @Quot.{u} α r → β 68 | return env.add <| .quotInfo { 69 | name := ``Quot.lift, kind := .lift, levelParams := [`u, `v] 70 | type := (← read).mkForall #[α, r, β, f] <| .arrow sanity <| .arrow quot_r β 71 | } 72 | let quotMk_a := mkApp3 (.const ``Quot.mk [u]) α r a 73 | withLocalDecl `β (.arrow quot_r .prop) .implicit fun β => do 74 | let all_quot := (← read).mkForall #[a] <| .app β quotMk_a 75 | withLocalDecl `q quot_r .implicit fun q => do 76 | -- constant Quot.ind.{u} {α : Sort u} {r : α → α → Prop} {β : @Quot.{u} α r → Prop} : 77 | -- (∀ a : α, β (@Quot.mk.{u} α r a)) → ∀ q : @Quot.{u} α r, β q */ 78 | let env := env.add <| .quotInfo { 79 | name := ``Quot.ind, kind := .ind, levelParams := [`u] 80 | type := (← read).mkForall #[α, r, β] <| 81 | .forallE `mk all_quot ((← read).mkForall #[q] <| .app β q) .default 82 | } 83 | return markQuotInit env 84 | 85 | def quotReduceRec [Monad m] (e : Expr) (whnf : Expr → m Expr) : m (Option Expr) := do 86 | let .const fn _ := e.getAppFn | return none 87 | let cont mkPos argPos := do 88 | let args := e.getAppArgs 89 | if h : mkPos < args.size then 90 | let mk ← whnf args[mkPos] 91 | if !mk.isAppOfArity ``Quot.mk 3 then return none 92 | let mut r := Expr.app args[argPos]! mk.appArg! 93 | let elimArity := mkPos + 1 94 | if elimArity < args.size then 95 | r := mkAppRange r elimArity args.size args 96 | return some r 97 | else return none 98 | if fn == ``Quot.lift then cont 5 3 99 | else if fn == ``Quot.ind then cont 4 3 100 | else return none 101 | -------------------------------------------------------------------------------- /Lean4Lean/Environment.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.TypeChecker 2 | import Lean4Lean.Quot 3 | import Lean4Lean.Inductive.Add 4 | import Lean4Lean.Primitive 5 | 6 | namespace Lean4Lean 7 | open Lean hiding Environment Exception 8 | open TypeChecker Kernel Environment 9 | 10 | open private Lean.Kernel.Environment.add from Lean.Environment 11 | 12 | def checkConstantVal (env : Environment) (v : ConstantVal) (allowPrimitive := false) : M Unit := do 13 | checkName env v.name allowPrimitive 14 | checkDuplicatedUnivParams v.levelParams 15 | checkNoMVarNoFVar env v.name v.type 16 | let sort ← checkType v.type 17 | _ ← ensureSort sort v.type 18 | 19 | def addAxiom (env : Environment) (v : AxiomVal) (check := true) : 20 | Except Exception Environment := do 21 | if check then 22 | _ ← (checkConstantVal env v.toConstantVal).run env 23 | (safety := if v.isUnsafe then .unsafe else .safe) 24 | return env.add (.axiomInfo v) 25 | 26 | def addDefinition (env : Environment) (v : DefinitionVal) (check := true) : 27 | Except Exception Environment := do 28 | if let .unsafe := v.safety then 29 | -- Meta definition can be recursive. 30 | -- So, we check the header, add, and then type check the body. 31 | if check then 32 | _ ← (checkConstantVal env v.toConstantVal).run env (safety := .unsafe) 33 | let env' := env.add (.defnInfo v) 34 | if check then 35 | checkNoMVarNoFVar env' v.name v.value 36 | M.run env' (safety := .unsafe) (lctx := {}) (lparams := v.levelParams) do 37 | let valType ← TypeChecker.checkType v.value 38 | if !(← isDefEq valType v.type) then 39 | throw <| .declTypeMismatch env' (.defnDecl v) valType 40 | return env' 41 | else 42 | if check then 43 | M.run env (safety := .safe) (lctx := {}) (lparams := v.levelParams) do 44 | checkConstantVal env v.toConstantVal (← checkPrimitiveDef v) 45 | let valType ← TypeChecker.checkType v.value 46 | if !(← isDefEq valType v.type) then 47 | throw <| .declTypeMismatch env (.defnDecl v) valType 48 | return env.add (.defnInfo v) 49 | 50 | def addTheorem (env : Environment) (v : TheoremVal) (check := true) : 51 | Except Exception Environment := do 52 | if check then 53 | -- TODO(Leo): we must add support for handling tasks here 54 | M.run env (safety := .safe) (lctx := {}) (lparams := v.levelParams) do 55 | if !(← isProp v.type) then 56 | throw <| .thmTypeIsNotProp env v.name v.type 57 | checkConstantVal env v.toConstantVal 58 | let valType ← TypeChecker.checkType v.value 59 | if !(← isDefEq valType v.type) then 60 | throw <| .declTypeMismatch env (.thmDecl v) valType 61 | return env.add (.thmInfo v) 62 | 63 | def addOpaque (env : Environment) (v : OpaqueVal) (check := true) : 64 | Except Exception Environment := do 65 | if check then 66 | M.run env (safety := .safe) (lctx := {}) (lparams := v.levelParams) do 67 | checkConstantVal env v.toConstantVal 68 | let valType ← TypeChecker.checkType v.value 69 | if !(← isDefEq valType v.type) then 70 | throw <| .declTypeMismatch env (.opaqueDecl v) valType 71 | return env.add (.opaqueInfo v) 72 | 73 | def addMutual (env : Environment) (vs : List DefinitionVal) (check := true) : 74 | Except Exception Environment := do 75 | let v₀ :: _ := vs | throw <| .other "invalid empty mutual definition" 76 | if let .safe := v₀.safety then 77 | throw <| .other "invalid mutual definition, declaration is not tagged as unsafe/partial" 78 | if check then 79 | M.run env (safety := v₀.safety) (lctx := {}) (lparams := v₀.levelParams) do 80 | for v in vs do 81 | if v.safety != v₀.safety then 82 | throw <| .other 83 | "invalid mutual definition, declarations must have the same safety annotation" 84 | checkConstantVal env v.toConstantVal 85 | let mut env' := env 86 | for v in vs do 87 | env' := env'.add (.defnInfo v) 88 | if check then 89 | M.run env' (safety := v₀.safety) (lctx := {}) (lparams := v₀.levelParams) do 90 | for v in vs do 91 | checkNoMVarNoFVar env' v.name v.value 92 | let valType ← TypeChecker.checkType v.value 93 | if !(← isDefEq valType v.type) then 94 | throw <| .declTypeMismatch env' (.mutualDefnDecl vs) valType 95 | return env' 96 | 97 | /-- Type check given declaration and add it to the environment -/ 98 | def addDecl (env : Environment) (decl : Declaration) (check := true) : 99 | Except Exception Environment := do 100 | match decl with 101 | | .axiomDecl v => addAxiom env v check 102 | | .defnDecl v => addDefinition env v check 103 | | .thmDecl v => addTheorem env v check 104 | | .opaqueDecl v => addOpaque env v check 105 | | .mutualDefnDecl v => addMutual env v check 106 | | .quotDecl => addQuot env 107 | | .inductDecl lparams nparams types isUnsafe => 108 | let allowPrimitive ← checkPrimitiveInductive env lparams nparams types isUnsafe 109 | addInductive env lparams nparams types isUnsafe allowPrimitive 110 | -------------------------------------------------------------------------------- /Lean4Lean/Environment/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean.Environment 2 | import Batteries.Tactic.OpenPrivate 3 | 4 | namespace Lean.Kernel.Environment 5 | 6 | def contains (env : Environment) (n : Name) : Bool := 7 | env.constants.contains n 8 | 9 | def get (env : Environment) (n : Name) : Except Exception ConstantInfo := 10 | match env.find? n with 11 | | some ci => pure ci 12 | | none => throw <| .unknownConstant env n 13 | 14 | def checkDuplicatedUnivParams : List Name → Except Exception Unit 15 | | [] => pure () 16 | | p :: ls => do 17 | if p ∈ ls then 18 | throw <| .other 19 | s!"failed to add declaration to environment, duplicate universe level parameter: '{p}'" 20 | checkDuplicatedUnivParams ls 21 | 22 | def checkNoMVar (env : Environment) (n : Name) (e : Expr) : Except Exception Unit := do 23 | if e.hasMVar then 24 | throw <| .declHasMVars env n e 25 | 26 | def checkNoFVar (env : Environment) (n : Name) (e : Expr) : Except Exception Unit := do 27 | if e.hasFVar then 28 | throw <| .declHasFVars env n e 29 | 30 | def checkNoMVarNoFVar (env : Environment) (n : Name) (e : Expr) : Except Exception Unit := do 31 | checkNoMVar env n e 32 | checkNoFVar env n e 33 | 34 | def primitives : NameSet := .ofList [ 35 | ``Bool, ``Bool.false, ``Bool.true, 36 | ``Nat, ``Nat.zero, ``Nat.succ, 37 | ``Nat.add, ``Nat.pred, ``Nat.sub, ``Nat.mul, ``Nat.pow, 38 | ``Nat.gcd, ``Nat.mod, ``Nat.div, ``Nat.beq, ``Nat.ble, 39 | ``Nat.bitwise, ``Nat.land, ``Nat.lor, ``Nat.xor, 40 | ``Nat.shiftLeft, ``Nat.shiftRight, 41 | ``String, ``String.mk] 42 | 43 | /-- 44 | Returns true iff `constName` is a non-recursive inductive datatype that has only one constructor and no indices. 45 | 46 | Such types have special kernel support. This must be in sync with `is_structure_like`. 47 | -/ 48 | def isStructureLike (env : Environment) (constName : Name) : Bool := 49 | match env.find? constName with 50 | | some (.inductInfo { isRec := false, ctors := [_], numIndices := 0, .. }) => true 51 | | _ => false 52 | 53 | def checkName (env : Environment) (n : Name) 54 | (allowPrimitive := false) : Except Exception Unit := do 55 | if env.contains n then 56 | throw <| .alreadyDeclared env n 57 | unless allowPrimitive do 58 | if primitives.contains n then 59 | throw <| .other s!"unexpected use of primitive name {n}" 60 | 61 | open private subsumesInfo Kernel.Environment.mk moduleNames moduleNameMap parts toEffectiveImport 62 | from Lean.Environment 63 | 64 | def empty (mainModule : Name) (trustLevel : UInt32 := 0) : Environment := 65 | Kernel.Environment.mk 66 | (constants := {}) 67 | (quotInit := false) 68 | (diagnostics := {}) 69 | (const2ModIdx := {}) 70 | (extensions := #[]) 71 | (irBaseExts := #[]) 72 | (header := { mainModule, trustLevel }) 73 | 74 | def throwAlreadyImported (s : ImportState) (const2ModIdx : Std.HashMap Name ModuleIdx) 75 | (modIdx : Nat) (cname : Name) : Except Exception α := do 76 | let modName := (moduleNames s)[modIdx]! 77 | let constModName := (moduleNames s)[const2ModIdx[cname]!.toNat]! 78 | throw <| .other 79 | s!"import {modName} failed, environment already contains '{cname}' from {constModName}" 80 | 81 | def finalizeImport (s : ImportState) (imports : Array Import) (mainModule : Name) 82 | (trustLevel : UInt32 := 0) : Except Exception Environment := do 83 | let modules := (moduleNames s).filterMap ((moduleNameMap s)[·]?) 84 | let moduleData ← modules.mapM fun mod => do 85 | let some data := mod.mainModule? | 86 | throw <| .other s!"missing data file for module {mod.module}" 87 | return data 88 | let numPrivateConsts := moduleData.foldl (init := 0) fun numPrivateConsts data => Id.run do 89 | numPrivateConsts + data.constants.size + data.extraConstNames.size 90 | let mut const2ModIdx := .emptyWithCapacity (capacity := numPrivateConsts) 91 | let mut privateConstantMap := .emptyWithCapacity (capacity := numPrivateConsts) 92 | for h : modIdx in [0:moduleData.size] do 93 | let data := moduleData[modIdx] 94 | for cname in data.constNames, cinfo in data.constants do 95 | match privateConstantMap.getThenInsertIfNew? cname cinfo with 96 | | (cinfoPrev?, constantMap') => 97 | privateConstantMap := constantMap' 98 | if let some cinfoPrev := cinfoPrev? then 99 | -- Recall that the map has not been modified when `cinfoPrev? = some _`. 100 | if subsumesInfo cinfo cinfoPrev then 101 | privateConstantMap := privateConstantMap.insert cname cinfo 102 | else if !subsumesInfo cinfoPrev cinfo then 103 | throwAlreadyImported s const2ModIdx modIdx cname 104 | const2ModIdx := const2ModIdx.insertIfNew cname modIdx 105 | for cname in data.extraConstNames do 106 | const2ModIdx := const2ModIdx.insertIfNew cname modIdx 107 | 108 | return Kernel.Environment.mk 109 | (constants := SMap.fromHashMap privateConstantMap false) 110 | (quotInit := !imports.isEmpty) -- We assume `Init.Prelude` initializes quotient module 111 | (diagnostics := {}) 112 | (const2ModIdx := const2ModIdx) 113 | (extensions := #[]) 114 | (irBaseExts := #[]) 115 | (header := { 116 | trustLevel, imports, moduleData, mainModule, isModule := false 117 | regions := modules.flatMap (parts · |>.map (·.2)) 118 | modules := modules.map toEffectiveImport 119 | }) 120 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Environment/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.LocalContext 2 | import Lean4Lean.Theory.Typing.EnvLemmas 3 | 4 | namespace Lean4Lean 5 | open Lean hiding Environment Exception 6 | open Kernel 7 | 8 | theorem ConstantInfo.hasValue_eq (ci : ConstantInfo) : ci.hasValue = ci.value?.isSome := by 9 | cases ci <;> rfl 10 | 11 | theorem ConstantInfo.value!_eq (ci : ConstantInfo) : ci.value! = ci.value?.get! := by 12 | cases ci <;> simp [ConstantInfo.value?, ConstantInfo.value!] 13 | 14 | def isAsSafeAs : DefinitionSafety → ConstantInfo → Bool 15 | | .unsafe, _ => true 16 | | .partial, ci => !ci.isUnsafe 17 | | .safe, ci => !ci.isUnsafe && !ci.isPartial 18 | 19 | theorem isAsSafeAs_of_safe : isAsSafeAs .safe ci → isAsSafeAs safety ci := by 20 | cases safety <;> simp +contextual [isAsSafeAs] 21 | 22 | variable (safety : DefinitionSafety) (env : VEnv) in 23 | def TrConstant (ci : ConstantInfo) (ci' : VConstant) : Prop := 24 | isAsSafeAs safety ci ∧ ci.levelParams.length = ci'.uvars ∧ 25 | TrExprS env ci.levelParams [] ci.type ci'.type 26 | 27 | variable (safety : DefinitionSafety) (env : VEnv) in 28 | def TrConstVal (ci : ConstantInfo) (ci' : VConstVal) : Prop := 29 | TrConstant safety env ci ci'.toVConstant ∧ ci.name = ci'.name 30 | 31 | variable (safety : DefinitionSafety) (env : VEnv) in 32 | def TrDefVal (ci : ConstantInfo) (ci' : VDefVal) : Prop := 33 | TrConstVal safety env ci ci'.toVConstVal ∧ 34 | TrExprS env ci.levelParams [] ci.value! ci'.value 35 | 36 | def AddQuot1 (name : Name) (kind : QuotKind) (ci' : VConstant) (P : ConstMap → VEnv → Prop) 37 | (m : ConstMap) (env : VEnv) : Prop := 38 | ∃ levelParams type env', 39 | let ci := .quotInfo { name, kind, levelParams, type } 40 | TrConstant .safe env ci ci' ∧ 41 | env.addConst name ci' = some env' ∧ 42 | P (m.insert name ci) env' 43 | 44 | theorem AddQuot1.to_addQuot 45 | (H1 : ∀ m env, P m env → f env = some env') 46 | (m env) (H : AddQuot1 name kind ci' P m env) : 47 | env.addConst name (some ci') >>= f = some env' := by 48 | let ⟨_, _, _, h1, h2, h3⟩ := H 49 | simpa using ⟨_, h2, H1 _ _ h3⟩ 50 | 51 | theorem AddQuot1.le 52 | (H1 : ∀ m env, P m env → env ≤ env₀) 53 | (m env) (H : AddQuot1 name kind ci' P m env) : env ≤ env₀ := 54 | let ⟨_, _, _, _, h2, h3⟩ := H 55 | .trans (VEnv.addConst_le h2) (H1 _ _ h3) 56 | 57 | def AddQuot (m₁ m₂ : ConstMap) (env₁ env₂ : VEnv) : Prop := 58 | AddQuot1 ``Quot .type quotConst (m := m₁) (env := env₁) <| 59 | AddQuot1 ``Quot.mk .ctor quotMkConst <| 60 | AddQuot1 ``Quot.lift .lift quotLiftConst <| 61 | AddQuot1 ``Quot.ind .ind quotIndConst (· = m₂ ∧ ·.addDefEq quotDefEq = env₂) 62 | 63 | nonrec theorem AddQuot.to_addQuot (H : AddQuot m₁ m₂ env₁ env₂) : env₁.addQuot = some env₂ := 64 | open AddQuot1 in (to_addQuot <| to_addQuot <| to_addQuot <| to_addQuot (by simp)) _ _ H 65 | 66 | nonrec theorem AddQuot.le (H : AddQuot m₁ m₂ env₁ env₂) : env₁ ≤ env₂ := 67 | open AddQuot1 in (le <| le <| le <| le fun _ _ h => h.2 ▸ VEnv.addDefEq_le) _ _ H 68 | 69 | inductive AddInduct (m₁ : ConstMap) (env₁ : VEnv) (decl : VInductDecl) 70 | (m₂ : ConstMap) (env₂ : VEnv) : Prop 71 | -- TODO 72 | 73 | nonrec theorem AddInduct.to_addInduct 74 | (H : AddInduct m₁ env₁ decl m₂ env₂) : env₁.addInduct decl = some env₂ := 75 | nomatch H 76 | 77 | variable (safety : DefinitionSafety) in 78 | inductive TrEnv' : ConstMap → Bool → VEnv → Prop where 79 | | empty : TrEnv' {} false .empty 80 | | block : 81 | ¬isAsSafeAs safety ci → 82 | env.addConst ci.name none = some env' → 83 | TrEnv' C Q env → 84 | TrEnv' (C.insert ci.name ci) Q env' 85 | | axiom : 86 | TrConstant safety env (.axiomInfo ci) ci' → 87 | ci'.WF env → 88 | env.addConst ci.name (some ci') = some env' → 89 | TrEnv' C Q env → 90 | TrEnv' (C.insert ci.name (.axiomInfo ci)) Q env' 91 | | defn {ci' : VDefVal} : 92 | TrDefVal safety env (.defnInfo ci) ci' → 93 | ci'.WF env → 94 | env.addConst ci.name (some ci'.toVConstant) = some env' → 95 | TrEnv' C Q env → 96 | TrEnv' (C.insert ci.name (.defnInfo ci)) Q (env'.addDefEq ci'.toDefEq) 97 | | opaque {ci' : VDefVal} : 98 | TrDefVal safety env (.opaqueInfo ci) ci' → 99 | ci'.WF env → 100 | env.addConst ci.name (some ci'.toVConstant) = some env' → 101 | TrEnv' C Q env → 102 | TrEnv' (C.insert ci.name (.opaqueInfo ci)) Q env' 103 | | quot : 104 | env.QuotReady → 105 | AddQuot C C' env env' → 106 | TrEnv' C false env → 107 | TrEnv' C' true env' 108 | | induct : 109 | decl.WF env → 110 | AddInduct C env decl C' env' → 111 | TrEnv' C Q env → 112 | TrEnv' C' Q env' 113 | 114 | def TrEnv (safety : DefinitionSafety) (env : Environment) (venv : VEnv) : Prop := 115 | TrEnv' safety env.constants env.quotInit venv 116 | 117 | theorem TrEnv'.wf (H : TrEnv' safety C Q venv) : venv.WF := by 118 | induction H with 119 | | empty => exact ⟨_, .empty⟩ 120 | | block _ h _ ih => 121 | have ⟨_, H⟩ := ih 122 | exact ⟨_, H.decl <| .block h⟩ 123 | | «axiom» _ h1 h2 _ ih => 124 | have ⟨_, H⟩ := ih 125 | exact ⟨_, H.decl <| .axiom (ci := ⟨_, _⟩) h1 h2⟩ 126 | | defn h1 h2 h3 _ ih => 127 | have ⟨_, H⟩ := ih 128 | have := h1.1.2; dsimp [ConstantInfo.name, ConstantInfo.toConstantVal] at this 129 | exact ⟨_, H.decl <| .def h2 (this ▸ h3)⟩ 130 | | «opaque» h1 h2 h3 _ ih => 131 | have ⟨_, H⟩ := ih 132 | have := h1.1.2; dsimp [ConstantInfo.name, ConstantInfo.toConstantVal] at this 133 | exact ⟨_, H.decl <| .opaque h2 (this ▸ h3)⟩ 134 | | quot h1 h2 _ ih => 135 | have ⟨_, H⟩ := ih 136 | exact ⟨_, H.decl <| .quot h1 h2.to_addQuot⟩ 137 | | induct h1 h2 _ ih => 138 | have ⟨_, H⟩ := ih 139 | exact ⟨_, H.decl <| .induct h1 h2.to_addInduct⟩ 140 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Lean-for-Lean 2 | 3 | This is an implementation of the Lean 4 kernel written in (mostly) pure Lean 4. 4 | It is derived directly from the C++ kernel implementation, and as such likely 5 | shares some implementation bugs with it (it's not really an independent 6 | implementation), although it also benefits from the same algorithmic performance 7 | improvements existing in the C++ Lean kernel. 8 | 9 | The project also houses some metatheory regarding the Lean 10 | system, in the same general direction as the 11 | [MetaCoq project](https://github.com/MetaCoq/metacoq/). 12 | 13 | ## Building 14 | 15 | To compile the code, you will need [Lean](https://lean-lang.org/lean4/doc/quickstart.html), or more specifically `elan`, the Lean version manager, which will make sure you have the right version of lean as specified in the [lean-toolchain](lean-toolchain) file. Assuming you have elan, the project can be compiled with: 16 | 17 | ``` 18 | lake build 19 | ``` 20 | 21 | This builds all components but you can also build them separately: 22 | 23 | * `lake build Lean4Lean` builds the `Lean4Lean` library interface, and does not include any of the proofs. 24 | * `lake build lean4lean` (note the capitalization!) builds the `lean4lean` command line tool, again without building proofs. 25 | * `lake build Lean4Lean.Theory` contains the Lean metatheory and properties. 26 | * `lake build Lean4Lean.Verify` is the proof that the `Lean4Lean` implementation satisfies the `Lean4Lean.Theory` abstract specification. 27 | 28 | ## Running 29 | 30 | After `lake build lean4lean`, the executable will be in `.lake/build/bin/lean4lean`. Because it requires some environment variables to be set for search paths which are provided by lake, you should evaluate it like `lake env .lake/build/bin/lean4lean`. 31 | 32 | If you run this as is (with no additional arguments), it will check every olean in the `lean4lean` package itself, which is probably not what you want. To check a different Lean package you should navigate the directory of the target project, then use `lake env path/to/lean4lean/.lake/build/bin/lean4lean ` to run `lean4lean` in the context of the target project. The command line arguments are: 33 | 34 | > `lean4lean [--fresh] [-v|--verbose] [--compare] [MOD]` 35 | 36 | * `MOD`: an optional lean module name, like `Lean4Lean.Verify`. If provided, the specified module will be checked (single-threaded); otherwise, all modules on the Lean search path will be checked (multithreaded). 37 | * `--fresh`: Only valid when a `MOD` is provided. In this mode, the module and all its imports will be rebuilt from scratch, checking all dependencies of the module. The behavior without the flag is to only check the module itself, assuming all imports are correct. 38 | * `--verbose`: shows the name of each declaration before adding it to the environment. Useful to know if the kernel got stuck on something. 39 | * `--compare`: If lean4lean takes more than a second on a given definition, we also check the C++ kernel performance to see if it is also slow on the same definition and report if lean4lean is abnormally slow in comparison. 40 | 41 | ## More documentation 42 | 43 | * [bugs-found.md](bugs-found.md): A list of kernel bugs that the lean4lean project has uncovered. 44 | * [divergences.md](divergences.md): A list of deliberate divergences between lean's kernel and the lean4lean kernel. 45 | 46 | ## (Selected) file breakdown 47 | 48 | * `Main.lean`: command line app 49 | * `Lean4Lean`: source files 50 | * `Environment.lean`: library entry point 51 | * `TypeChecker.lean`: main recursive function 52 | * `Inductive` 53 | * `Add.lean`: constructing inductive recursors 54 | * `Reduce.lean`: inductive iota rules 55 | * `Quot.lean`: quotient types handling 56 | * `Primitive.lean`: checking correctness of built-ins 57 | * `Std`: stuff that should exist upstream 58 | * `Theory`: lean metatheory 59 | * `VLevel.lean`: level expressions 60 | * `VExpr.lean`: expressions (boring de Bruijn variable theorems are here) 61 | * `VDecl.lean`: declarations 62 | * `VEnv.lean`: environment 63 | * `Meta.lean`: elaborator producing `VExpr`s 64 | * `Inductive.lean`: inductive types 65 | * `Quot.lean`: quotient types 66 | * `Typing` 67 | * `Basic.lean`: The typing relation itself 68 | * `Lemmas.lean`: theorems about the typing relation 69 | * `Meta.lean`: tactic for proving typing judgments 70 | * `Strong.lean`: proof that you can have all the inductive hypotheses 71 | * `UniqueTyping.lean`: conjectures about the typing relation 72 | * `Env.lean`: typing for environments 73 | * `Verify`: relation between the metatheory and the kernel 74 | * `Axioms.lean`: theorems about upstream opaques that shouldn't be opaque 75 | * `Expr.lean`: correctness of basics on `Expr` 76 | * `Level.lean`: correctness of basics on `Level` 77 | * `VLCtx.lean`: a "translation context" suitable for translating expressions 78 | * `LocalContext.lean`: properties of lean's `LocalContext` type 79 | * `NameGenerator.lean`: properties of the fresh name generator 80 | * `Typing` 81 | * `Expr.lean`: translating expressions (`TrExpr` is here) 82 | * `Lemmas.lean`: properties of `TrExpr` 83 | * `ConditionallyTyped.lean`: properties of expressions in caches that may be out of scope 84 | * `Environment` 85 | * `Basic.lean`: translating environments 86 | * `Lemmas.lean`: properties of `TrEnv` 87 | * `TypeChecker` 88 | * `Basic.lean`: typechecker invariants 89 | * `EquivManager.lean`: invariants for the union-find defeq cache 90 | * `InferType.lean`: correctness of `inferType` 91 | * `WHNF.lean`: correctness of `whnf` 92 | * `IsDefEq.lean`: correctness of `isDefEq` 93 | * `TypeChecker.lean`: top-level typechecker correctness 94 | * `Experimental`: work in progress formalizations and ideas 95 | * `Stratified.lean`: stratified typing judgment 96 | * `StratifiedUntyped.lean` another stratified typing judgment 97 | * `ParallelReduction.lean`: stuff related to church-rosser 98 | * `Stronger.lean`: a more heavily annotated typing judgment 99 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/VLevel.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Std.Basic 2 | 3 | namespace Lean4Lean 4 | 5 | export Lean (Name) 6 | 7 | inductive VLevel where 8 | | zero : VLevel 9 | | succ : VLevel → VLevel 10 | | max : VLevel → VLevel → VLevel 11 | | imax : VLevel → VLevel → VLevel 12 | | param : Nat → VLevel 13 | 14 | namespace VLevel 15 | 16 | variable (n : Nat) in 17 | def WF : VLevel → Prop 18 | | .zero => True 19 | | .succ l => l.WF 20 | | .max l₁ l₂ => l₁.WF ∧ l₂.WF 21 | | .imax l₁ l₂ => l₁.WF ∧ l₂.WF 22 | | .param i => i < n 23 | 24 | instance decidable_WF : ∀ {l}, Decidable (WF n l) 25 | | .zero => instDecidableTrue 26 | | .succ l => @decidable_WF _ l 27 | | .max .. | .imax .. => @instDecidableAnd _ _ decidable_WF decidable_WF 28 | | .param _ => Nat.decLt .. 29 | 30 | variable (ls : List Nat) in 31 | def eval : VLevel → Nat 32 | | .zero => 0 33 | | .succ l => l.eval + 1 34 | | .max l₁ l₂ => l₁.eval.max l₂.eval 35 | | .imax l₁ l₂ => l₁.eval.imax l₂.eval 36 | | .param i => ls.getD i 0 37 | 38 | protected def LE (a b : VLevel) : Prop := ∀ ls, a.eval ls ≤ b.eval ls 39 | 40 | instance : LE VLevel := ⟨VLevel.LE⟩ 41 | 42 | theorem le_refl (a : VLevel) : a ≤ a := fun _ => Nat.le_refl _ 43 | theorem le_trans {a b c : VLevel} (h1 : a ≤ b) (h2 : b ≤ c) : a ≤ c := 44 | fun _ => Nat.le_trans (h1 _) (h2 _) 45 | 46 | theorem zero_le : zero ≤ a := fun _ => Nat.zero_le _ 47 | 48 | theorem le_succ : a ≤ succ a := fun _ => Nat.le_succ _ 49 | 50 | theorem succ_le_succ (h : a ≤ b) : succ a ≤ succ b := fun _ => Nat.succ_le_succ (h _) 51 | 52 | theorem le_max_left : a ≤ max a b := fun _ => Nat.le_max_left .. 53 | theorem le_max_right : b ≤ max a b := fun _ => Nat.le_max_right .. 54 | 55 | protected def Equiv (a b : VLevel) : Prop := a.eval = b.eval 56 | 57 | instance : HasEquiv VLevel := ⟨VLevel.Equiv⟩ 58 | 59 | theorem equiv_def' {a b : VLevel} : a ≈ b ↔ a.eval = b.eval := .rfl 60 | theorem equiv_def {a b : VLevel} : a ≈ b ↔ ∀ ls, a.eval ls = b.eval ls := funext_iff 61 | 62 | theorem equiv_congr_left {a b c : VLevel} (h : a ≈ b) : a ≈ c ↔ b ≈ c := 63 | iff_of_eq (congrArg (· = _) h) 64 | 65 | theorem equiv_congr_right {a b c : VLevel} (h : a ≈ b) : c ≈ a ↔ c ≈ b := 66 | iff_of_eq (congrArg (_ = ·) h) 67 | 68 | theorem le_antisymm_iff {a b : VLevel} : a ≈ b ↔ a ≤ b ∧ b ≤ a := 69 | equiv_def.trans <| (forall_congr' fun _ => Nat.le_antisymm_iff).trans forall_and 70 | 71 | theorem succ_congr {a b : VLevel} (h : a ≈ b) : succ a ≈ succ b := by 72 | simpa [equiv_def, eval] using h 73 | 74 | theorem succ_congr_iff {a b : VLevel} : succ a ≈ succ b ↔ a ≈ b := by 75 | simp [equiv_def, eval] 76 | 77 | theorem max_congr (h₁ : a₁ ≈ b₁) (h₂ : a₂ ≈ b₂) : max a₁ a₂ ≈ max b₁ b₂ := by 78 | simp_all [equiv_def, eval] 79 | 80 | theorem imax_congr (h₁ : a₁ ≈ b₁) (h₂ : a₂ ≈ b₂) : imax a₁ a₂ ≈ imax b₁ b₂ := by 81 | simp_all [equiv_def, eval] 82 | 83 | theorem max_comm : max a b ≈ max b a := by simp [equiv_def, eval, Nat.max_comm] 84 | 85 | theorem LE.max_eq_left (h : b.LE a) : max a b ≈ a := by 86 | simp [equiv_def, eval, Nat.max_eq_left (h _)] 87 | 88 | theorem LE.max_eq_right (h : a.LE b) : max a b ≈ b := by 89 | simp [equiv_def, eval, Nat.max_eq_right (h _)] 90 | 91 | theorem max_self : max a a ≈ a := by simp [equiv_def, eval] 92 | 93 | theorem zero_imax : imax zero a ≈ a := by 94 | simp [equiv_def, eval, Nat.imax, eq_comm (b := 0)] 95 | 96 | theorem imax_self : imax a a ≈ a := by 97 | simp [equiv_def, eval, Nat.imax, eq_comm (b := 0)] 98 | 99 | def IsNeverZero (a : VLevel) : Prop := ∀ ls, a.eval ls ≠ 0 100 | 101 | theorem IsNeverZero.imax_eq_max (h : IsNeverZero b) : imax a b ≈ max a b := by 102 | simp_all [equiv_def, eval, Nat.imax, IsNeverZero] 103 | 104 | variable (ls : List VLevel) in 105 | def inst : VLevel → VLevel 106 | | .zero => .zero 107 | | .succ l => .succ l.inst 108 | | .max l₁ l₂ => .max l₁.inst l₂.inst 109 | | .imax l₁ l₂ => .imax l₁.inst l₂.inst 110 | | .param i => ls.getD i .zero 111 | 112 | theorem inst_inst {l : VLevel} : (l.inst ls).inst ls' = l.inst (ls.map (inst ls')) := by 113 | induction l <;> simp [inst, *, List.getD_eq_getElem?_getD, List.getElem?_map] 114 | case param n => cases ls[n]? <;> simp [inst] 115 | 116 | def params (n : Nat) : List VLevel := (List.range n).map .param 117 | 118 | @[simp] theorem params_length {n : Nat} : (params n).length = n := by simp [params] 119 | 120 | theorem params_wf {n : Nat} : ∀ ⦃l⦄, l ∈ params n → l.WF n := by simp [params, WF] 121 | 122 | theorem inst_id {l : VLevel} (h : l.WF u) : l.inst (params u) = l := by 123 | induction l <;> simp_all [params, inst, WF, List.getD_eq_getElem?_getD] 124 | 125 | theorem inst_map_id (h : ls.length = n) : (params n).map (inst ls) = ls := by 126 | subst n; simp [params]; apply List.ext_get (by simp) 127 | intro i _ _; simp [inst]; rw [List.getElem?_eq_getElem]; rfl 128 | 129 | theorem eval_inst {l : VLevel} : (l.inst ls).eval ns = l.eval (ls.map (eval ns)) := by 130 | induction l <;> simp [eval, inst, *, List.getD_eq_getElem?_getD] 131 | case param n => cases ls[n]? <;> simp [eval] 132 | 133 | theorem WF.inst {l : VLevel} (H : ∀ l ∈ ls, l.WF n) : (l.inst ls).WF n := by 134 | induction l with 135 | | zero => trivial 136 | | succ _ ih => exact ih 137 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => exact ⟨ih1, ih2⟩ 138 | | param i => 139 | simp [VLevel.inst, List.getD_eq_getElem?_getD] 140 | cases e : ls[i]? with 141 | | none => trivial 142 | | some => exact H _ (List.mem_of_getElem? e) 143 | 144 | theorem id_WF : ∀ l ∈ (List.range u).map param, l.WF u := by simp [WF] 145 | 146 | theorem inst_congr {l : VLevel} (h1 : l ≈ l') (h2 : List.Forall₂ (·≈·) ls ls') : 147 | l.inst ls ≈ l'.inst ls' := by 148 | simp [equiv_def, eval_inst, ← equiv_def.1 h1] 149 | intro ns; congr 1 150 | induction h2 with 151 | | nil => rfl 152 | | cons h2 => simp [*, equiv_def.1 h2] 153 | 154 | theorem inst_congr_l {l : VLevel} (h1 : l ≈ l') : l.inst ls ≈ l'.inst ls := 155 | inst_congr h1 <| .rfl fun _ _ => rfl 156 | 157 | variable (ls : List Name) in 158 | def ofLevel : Lean.Level → Option VLevel 159 | | .zero => return .zero 160 | | .succ l => return .succ (← ofLevel l) 161 | | .max l₁ l₂ => return .max (← ofLevel l₁) (← ofLevel l₂) 162 | | .imax l₁ l₂ => return .imax (← ofLevel l₁) (← ofLevel l₂) 163 | | .param n => 164 | let i := ls.idxOf n 165 | if i < ls.length then some (.param i) else none 166 | | .mvar _ => none 167 | 168 | theorem WF.of_ofLevel (h : ofLevel ls l = some l') : l'.WF ls.length := by 169 | induction l generalizing l' with simp [ofLevel, bind] at h 170 | | zero => cases h; trivial 171 | | succ _ ih => obtain ⟨l', h, ⟨⟩⟩ := h; exact @ih l' h 172 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => obtain ⟨_, h1, _, h2, ⟨⟩⟩ := h; exact ⟨ih1 h1, ih2 h2⟩ 173 | | param n => exact h.2 ▸ h.1 174 | 175 | theorem WF.of_mapM_ofLevel (h : List.mapM (VLevel.ofLevel Us) us = some us') 176 | (a) (hl : a ∈ us') : VLevel.WF Us.length a := by 177 | rw [List.mapM_eq_some] at h 178 | have ⟨_, _, h⟩ := h.forall_exists_r _ hl; exact .of_ofLevel h 179 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Level.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.VLevel 2 | import Lean4Lean.Level 3 | import Lean4Lean.Verify.Axioms 4 | import Std.Tactic.BVDecide 5 | 6 | namespace Lean.Level 7 | open Lean4Lean 8 | 9 | attribute [simp] mkLevelSucc mkLevelMax mkLevelIMax updateSucc! updateMax! updateIMax! 10 | 11 | -- variable (ls : List Name) in 12 | -- def _root_.Lean4Lean.VLevel.toLevel : VLevel → Level 13 | -- | .zero => .zero 14 | -- | .succ l => .succ l.toLevel 15 | -- | .max l₁ l₂ => .max l₁.toLevel l₂.toLevel 16 | -- | .imax l₁ l₂ => .imax l₁.toLevel l₂.toLevel 17 | -- | .param n => match ls.get? n with 18 | -- | some l => .param l 19 | -- | none => .zero 20 | 21 | -- theorem toLevel_inj {ls : List Name} (d : ls.Nodup) 22 | -- {l₁ l₂ : VLevel} (eq : l₁.toLevel ls = l₂.toLevel ls) : l₁ = l₂ := sorry 23 | 24 | @[simp] def getOffset' : Level → Nat 25 | | succ u => getOffset' u + 1 26 | | _ => 0 27 | 28 | @[simp] theorem getOffset_eq (u : Level) : u.getOffset = u.getOffset' := go _ 0 where 29 | go (u : Level) (i) : u.getOffsetAux i = u.getOffset' + i := by 30 | unfold getOffsetAux getOffset'; split <;> simp 31 | rw [go]; simp [Nat.add_right_comm, Nat.add_assoc] 32 | 33 | theorem mkData_depth (H : d < 2 ^ 24) : (mkData h d hmv hp).depth.toNat = d := by 34 | rw [mkData_eq, mkData', if_neg (Nat.not_lt.2 (Nat.le_sub_one_of_lt H)), Data.depth] 35 | have : d.toUInt64.toUInt32.toNat = d := by simp; omega 36 | refine .trans ?_ this; congr 2 37 | rw [← UInt64.toBitVec_inj] 38 | have : d.toUInt64.toNat = d := by simp; omega 39 | have : d.toUInt64.toBitVec ≤ 0xffffff#64 := (this ▸ Nat.le_sub_one_of_lt H :) 40 | have : h.toUInt32.toUInt64.toBitVec ≤ 0xffffffff#64 := Nat.le_of_lt_succ h.toUInt32.1.1.2 41 | have hb : ∀ (b : Bool), b.toUInt64.toBitVec ≤ 1#64 := by decide 42 | have := hb hmv; have := hb hp 43 | change ( 44 | h.toUInt32.toUInt64.toBitVec + 45 | hmv.toUInt64.toBitVec <<< 32#64 + 46 | hp.toUInt64.toBitVec <<< 33#64 + 47 | d.toUInt64.toBitVec <<< 40#64) >>> 40#64 = d.toUInt64.toBitVec 48 | bv_decide 49 | 50 | theorem mkData_hasParam (H : d < 2 ^ 24) : (mkData h d hmv hp).hasParam = hp := by 51 | rw [mkData_eq, mkData', if_neg (Nat.not_lt.2 (Nat.le_sub_one_of_lt H))] 52 | simp [Data.hasParam, (· == ·), ← UInt64.toBitVec_inj] 53 | have : h.toUInt32.toUInt64.toBitVec ≤ 0xffffffff#64 := Nat.le_of_lt_succ h.toUInt32.1.1.2 54 | have hb : ∀ (b : Bool), b.toUInt64.toBitVec ≤ 1#64 := by decide 55 | have := hb hmv; have := hb hp 56 | let L := (( 57 | h.toUInt32.toUInt64.toBitVec + 58 | hmv.toUInt64.toBitVec <<< 32#64 + 59 | hp.toUInt64.toBitVec <<< 33#64 + 60 | d.toUInt64.toBitVec <<< 40#64) >>> 33#64) &&& 1#64 61 | change decide (L = 1#64) = hp 62 | rw [show L = hp.toUInt64.toBitVec by bv_decide] 63 | cases hp <;> decide 64 | 65 | theorem mkData_hasMVar (H : d < 2 ^ 24) : (mkData h d hmv hp).hasMVar = hmv := by 66 | rw [mkData_eq, mkData', if_neg (Nat.not_lt.2 (Nat.le_sub_one_of_lt H))] 67 | simp [Data.hasMVar, (· == ·), ← UInt64.toBitVec_inj] 68 | have : h.toUInt32.toUInt64.toBitVec ≤ 0xffffffff#64 := Nat.le_of_lt_succ h.toUInt32.1.1.2 69 | have hb : ∀ (b : Bool), b.toUInt64.toBitVec ≤ 1#64 := by decide 70 | have := hb hmv; have := hb hp 71 | let L := (( 72 | h.toUInt32.toUInt64.toBitVec + 73 | hmv.toUInt64.toBitVec <<< 32#64 + 74 | hp.toUInt64.toBitVec <<< 33#64 + 75 | d.toUInt64.toBitVec <<< 40#64) >>> 32#64) &&& 1#64 76 | change decide (L = 1#64) = hmv 77 | rw [show L = hmv.toUInt64.toBitVec by bv_decide] 78 | cases hmv <;> decide 79 | 80 | theorem ofLevel_of_not_hasParam (Us) {l : Level} 81 | (hl : l.hasParam' = false) (hmv : l.hasMVar' = false) : 82 | ∃ u', VLevel.ofLevel Us l = some u' := by 83 | induction l <;> simp_all [hasParam', hasMVar', VLevel.ofLevel, exists_comm] 84 | 85 | def getUndefParam.F (ps : List Name) (l : Level) : StateT (Option Name) Id Bool := do 86 | if !l.hasParam || (← get).isSome then 87 | return false 88 | if let .param n := l then 89 | if n ∉ ps then 90 | set (some n) 91 | return true 92 | 93 | theorem getUndefParam_none {l : Level} (hmv : l.hasMVar' = false) : 94 | l.getUndefParam Us = none → ∃ u', VLevel.ofLevel Us l = some u' := by 95 | suffices ∀ s, ((l.forEach (getUndefParam.F Us)).run s).run.snd = none → s = none ∧ _ from 96 | (this _ · |>.2) 97 | have {l} (hmv : l.hasMVar' = false) 98 | {g} (H : ∀ {s'}, (g.run s').run.snd = none → s' = none ∧ 99 | (((getUndefParam.F Us l).run none).run = (true, none) → 100 | ∃ u', VLevel.ofLevel Us l = some u')) (s) : 101 | ((do if (!(← getUndefParam.F Us l)) = true then pure PUnit.unit else g) 102 | |>.run s).run.snd = none → 103 | s = none ∧ ∃ u', VLevel.ofLevel Us l = some u' := by 104 | simp; split <;> rename_i h 105 | · simp; revert h 106 | simp [getUndefParam.F]; split <;> [simp; split <;> [split <;> simp; simp]] 107 | rintro rfl; simp at * 108 | exact ofLevel_of_not_hasParam Us ‹_› hmv 109 | · refine fun h' => let ⟨h1, h2⟩ := H h'; have := ?_; ⟨this, h2 ?_⟩ 110 | · revert h h1 111 | simp [getUndefParam.F]; split <;> [simp; split <;> [split <;> simp; simp]] 112 | · revert h h1; subst s 113 | cases (getUndefParam.F Us l).run none; simp; rintro rfl rfl; rfl 114 | have lt {n a} : n + 1 < a → n < a := by omega 115 | induction l with ( 116 | refine this hmv fun h => ?_; clear this 117 | simp [hasMVar', VLevel.ofLevel, *] at *) 118 | | succ _ ih => 119 | have ⟨h, _, h1⟩ := ih hmv _ h 120 | exact ⟨h, fun _ => ⟨_, _, h1, rfl⟩⟩ 121 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => 122 | have ⟨h, _, h2⟩ := ih2 hmv.2 _ h 123 | have ⟨h, _, h1⟩ := ih1 hmv.1 _ h 124 | exact ⟨h, fun _ => ⟨_, _, h1, _, h2, rfl⟩⟩ 125 | | param => 126 | simp [getUndefParam.F, hasParam', List.idxOf_lt_length_iff, *] 127 | split <;> simp [*] 128 | | _ => simp [*] 129 | 130 | variable (s : Name → Level) in 131 | def substParams' (red : Bool) : Level → Level 132 | | .zero => .zero 133 | | .succ v => .succ (substParams' (v.hasParam ∧ red) v) 134 | | .max v₁ v₂ => 135 | let red := (v₁.hasParam ∨ v₂.hasParam) ∧ red 136 | (if red then mkLevelMax' else .max) (substParams' red v₁) (substParams' red v₂) 137 | | .imax v₁ v₂ => 138 | let red := (v₁.hasParam ∨ v₂.hasParam) ∧ red 139 | (if red then mkLevelIMax' else .imax) (substParams' red v₁) (substParams' red v₂) 140 | | .param n => s n 141 | | u => u 142 | 143 | theorem substParams_eq_self {u : Level} (h : u.hasParam' = false) : 144 | substParams' s red u = u := by 145 | induction u generalizing red <;> simp_all [substParams', hasParam'] 146 | 147 | open private substParams.go from Lean.Level in 148 | @[simp] theorem substParams_eq (u : Level) (s : Name → Option Level) : 149 | substParams u s = substParams' (fun x => (s x).getD (.param x)) true u := by 150 | unfold substParams 151 | induction u <;> simp [substParams.go, substParams', hasParam', ← Bool.or_eq_true] <;> 152 | split <;> simp [*, substParams_eq_self] <;> simp_all [substParams_eq_self] 153 | 154 | theorem substParams_id {u : Level} : 155 | substParams' .param false u = u := by induction u <;> simp_all [substParams'] 156 | 157 | theorem isEquiv'_wf (h : isEquiv' u v) 158 | (hu : VLevel.ofLevel ls u = some u') (hv : VLevel.ofLevel ls v = some v') : u' ≈ v' := sorry 159 | 160 | theorem isEquivList_wf (H : Level.isEquivList us vs) : 161 | List.mapM (VLevel.ofLevel Us) us = some us' → 162 | List.mapM (VLevel.ofLevel Us) vs = some vs' → us'.Forall₂ (· ≈ ·) vs' := by 163 | simp [Level.isEquivList] at H; revert us' vs' 164 | induction us generalizing vs with cases vs <;> simp [List.all2] at H <;> simp | cons u us ih 165 | rename_i v vs; rintro _ _ u' hu us' hus rfl v' hv vs' hvs rfl 166 | exact .cons (isEquiv'_wf H.1 hu hv) (ih H.2 hus hvs) 167 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Typing/ConditionallyTyped.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.Typing.Lemmas 2 | 3 | namespace Lean4Lean 4 | open VEnv Lean 5 | 6 | def ConditionallyTyped 7 | (ngen : NameGenerator) (env : VEnv) (Us : List Name) (Δ : VLCtx) (e : Expr) : Prop := 8 | Closed e ∧ FVarsIn ngen.Reserves e ∧ (FVarsIn (· ∈ Δ.fvars) e → ∃ e', TrExprS env Us Δ e e') 9 | 10 | theorem ConditionallyTyped.mk {Δ : VLCtx} 11 | (noBV : Δ.NoBV) (r : FVarsIn ngen.Reserves e) (H : TrExprS env Us Δ e e') : 12 | ConditionallyTyped ngen env Us Δ e := ⟨noBV ▸ H.closed, r, fun _ => ⟨_, H⟩⟩ 13 | 14 | theorem ConditionallyTyped.mono (H : ngen₁ ≤ ngen₂) : 15 | ConditionallyTyped ngen₁ env Us Δ e → ConditionallyTyped ngen₂ env Us Δ e 16 | | ⟨h1, h2, h3⟩ => ⟨h1, h2.mono fun _ h => h.mono H, h3⟩ 17 | 18 | theorem ConditionallyTyped.weakN_inv 19 | (henv : VEnv.WF env) (hΔ : VLCtx.WF env Us.length ((some fv, d) :: Δ)) 20 | (H : ConditionallyTyped ngen env Us ((some fv, d) :: Δ) e) : 21 | ConditionallyTyped ngen env Us Δ e := by 22 | refine ⟨H.1, H.2.1, fun H2 => ?_⟩ 23 | have := H.2.2 24 | have ⟨e', h⟩ := H.2.2 H2.fvars_cons 25 | exact TrExprS.weakFV_inv henv (.skip_fvar _ _ .refl) (.refl henv hΔ) h H.1 H2 26 | 27 | theorem ConditionallyTyped.fresh 28 | (henv : Ordered env) (hΔ : VLCtx.WF env Us.length ((some (⟨ngen.curr⟩, deps), d) :: Δ)) 29 | (H : ConditionallyTyped ngen env Us Δ e) : 30 | ConditionallyTyped ngen env Us ((some (⟨ngen.curr⟩, deps), d) :: Δ) e := by 31 | refine have ⟨H1, H2, H3⟩ := H; ⟨H1, H2, fun H4 => ?_⟩ 32 | refine have ⟨_, h⟩ := H3 (H4.mp ?_ H2); ⟨_, h.weakFV henv (.skip_fvar _ _ .refl) hΔ⟩ 33 | intro _ h1 h2; simp at h1; rcases h1 with rfl | h1 34 | · cases Nat.lt_irrefl _ (h2 _ rfl) 35 | · exact h1 36 | 37 | def ConditionallyHasType 38 | (ngen : NameGenerator) (env : VEnv) (Us : List Name) (Δ : VLCtx) (e A : Expr) : Prop := 39 | Closed e ∧ FVarsIn ngen.Reserves e ∧ Closed A ∧ FVarsIn ngen.Reserves A ∧ 40 | (FVarsIn (· ∈ Δ.fvars) e → ∃ e' A', TrTyping env Us Δ e A e' A') 41 | 42 | theorem ConditionallyHasType.mk {Δ : VLCtx} 43 | (noBV : Δ.NoBV) (he : TrTyping env Us Δ e A e' A') 44 | (re : FVarsIn ngen.Reserves e) (rA : FVarsIn ngen.Reserves A) : 45 | ConditionallyHasType ngen env Us Δ e A := by 46 | refine ⟨noBV ▸ he.2.1.closed, re, noBV ▸ he.2.2.1.closed, rA, fun _ => ⟨_, _, he.1, he.2⟩⟩ 47 | 48 | theorem ConditionallyHasType.typed : 49 | ConditionallyHasType ngen env Us Δ e A → ConditionallyTyped ngen env Us Δ e 50 | | ⟨c1, f1, _, _, H⟩ => ⟨c1, f1, fun h => let ⟨_, _, _, h, _⟩ := H h; ⟨_, h⟩⟩ 51 | 52 | theorem ConditionallyHasType.mono (H : ngen₁ ≤ ngen₂) : 53 | ConditionallyHasType ngen₁ env Us Δ e A → ConditionallyHasType ngen₂ env Us Δ e A 54 | | ⟨c1, f1, c2, f2, h'⟩ => ⟨c1, f1.mono fun _ h => h.mono H, c2, f2.mono fun _ h => h.mono H, h'⟩ 55 | 56 | theorem ConditionallyHasType.weakN_inv 57 | (henv : VEnv.WF env) (hΔ : VLCtx.WF env Us.length ((some fv, d) :: Δ)) 58 | (H : ConditionallyHasType ngen env Us ((some fv, d) :: Δ) e A) : 59 | ConditionallyHasType ngen env Us Δ e A := by 60 | have ⟨c1, f1, c2, f2, H⟩ := H 61 | refine ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 62 | have ⟨e', A', h1, h2, h3, h4⟩ := H H4.fvars_cons 63 | have W : VLCtx.FVLift Δ ((some fv, d) :: Δ) 0 (0 + d.depth) 0 := .skip_fvar _ _ .refl 64 | have ⟨e'', he⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h2 c1 H4 65 | have ee := h2.uniq henv (.refl henv hΔ) <| he.weakFV henv W hΔ 66 | have := hΔ.2.1 _ _ rfl 67 | have : IsFVarUpSet (· ∈ VLCtx.fvars Δ) ((some fv, d) :: Δ) := 68 | ⟨.fvars hΔ.1.fvwf, (hΔ.2.1 _ _ rfl).1.elim⟩ 69 | have ⟨_, hA⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h3 c2 <| h1 _ this H4 70 | have AA := h3.uniq henv (.refl henv hΔ) <| hA.weakFV henv W hΔ 71 | have h4 := h4.defeqU_r henv hΔ.toCtx AA |>.defeqU_l henv hΔ.toCtx ee 72 | have h4 := (HasType.weakN_iff henv hΔ.toCtx W.toCtx).1 h4 73 | refine ⟨_, _, fun P hP he' => ?_, he, hA, h4⟩ 74 | exact h1 _ 75 | ⟨(IsFVarUpSet.and_fvars hΔ.1.fvwf).1 hP, fun h => (hΔ.2.1 _ _ rfl).1.elim h.2⟩ 76 | (he'.mp (fun _ => .intro) he.fvarsIn) |>.mono fun _ => (·.1) 77 | 78 | theorem ConditionallyHasType.fresh 79 | (henv : Ordered env) 80 | (hΔ : VLCtx.WF env Us.length ((some (⟨ngen.curr⟩, deps), d) :: Δ)) 81 | (H : ConditionallyHasType ngen env Us Δ e A) : 82 | ConditionallyHasType ngen env Us ((some (⟨ngen.curr⟩, deps), d) :: Δ) e A := by 83 | refine have ⟨c1, f1, c2, f2, H⟩ := H; ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 84 | have ⟨_, _, h1, h2, h3, h4⟩ := H (H4.mp ?_ f1) 85 | · have W : VLCtx.FVLift Δ ((some (⟨ngen.curr⟩, deps), d) :: Δ) 0 (0 + d.depth) 0 := 86 | .skip_fvar _ _ .refl 87 | exact ⟨_, _, fun P hP => h1 _ hP.1, 88 | h2.weakFV henv W hΔ, h3.weakFV henv W hΔ, h4.weakN henv W.toCtx⟩ 89 | · intro _ h1 h2; simp at h1; rcases h1 with rfl | h1 90 | · cases Nat.lt_irrefl _ (h2 _ rfl) 91 | · exact h1 92 | 93 | def ConditionallyWHNF 94 | (ngen : NameGenerator) (env : VEnv) (Us : List Name) (Δ : VLCtx) (e e₁ : Expr) : Prop := 95 | Closed e ∧ FVarsIn ngen.Reserves e ∧ Closed e₁ ∧ FVarsIn ngen.Reserves e₁ ∧ 96 | (FVarsIn (· ∈ Δ.fvars) e → ∃ e', 97 | FVarsBelow Δ e e₁ ∧ TrExprS env Us Δ e e' ∧ TrExpr env Us Δ e₁ e') 98 | 99 | theorem ConditionallyWHNF.mk {Δ : VLCtx} 100 | (noBV : Δ.NoBV) (hb : FVarsBelow Δ e e₁) 101 | (he : TrExprS env Us Δ e e') (he₁ : TrExpr env Us Δ e₁ e') 102 | (re : FVarsIn ngen.Reserves e) (re₁ : FVarsIn ngen.Reserves e₁) : 103 | ConditionallyWHNF ngen env Us Δ e e₁ := by 104 | refine ⟨noBV ▸ he.closed, re, noBV ▸ he₁.closed, re₁, fun _ => ⟨_, hb, he, he₁⟩⟩ 105 | 106 | theorem ConditionallyWHNF.mono (H : ngen₁ ≤ ngen₂) : 107 | ConditionallyWHNF ngen₁ env Us Δ e A → ConditionallyWHNF ngen₂ env Us Δ e A 108 | | ⟨c1, f1, c2, f2, h'⟩ => ⟨c1, f1.mono fun _ h => h.mono H, c2, f2.mono fun _ h => h.mono H, h'⟩ 109 | 110 | theorem ConditionallyWHNF.weakN_inv 111 | (henv : VEnv.WF env) (hΔ : VLCtx.WF env Us.length ((some fv, d) :: Δ)) 112 | (H : ConditionallyWHNF ngen env Us ((some fv, d) :: Δ) e e₁) : 113 | ConditionallyWHNF ngen env Us Δ e e₁ := by 114 | have ⟨c1, f1, c2, f2, H⟩ := H 115 | refine ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 116 | have ⟨e', h1, h2, _, h3, h4⟩ := H H4.fvars_cons 117 | have W : VLCtx.FVLift Δ ((some fv, d) :: Δ) 0 (0 + d.depth) 0 := .skip_fvar _ _ .refl 118 | have ⟨e'', he⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h2 c1 H4 119 | have ee := h2.uniq henv (.refl henv hΔ) <| he.weakFV henv W hΔ 120 | have := hΔ.2.1 _ _ rfl 121 | have : IsFVarUpSet (· ∈ VLCtx.fvars Δ) ((some fv, d) :: Δ) := 122 | ⟨.fvars hΔ.1.fvwf, (hΔ.2.1 _ _ rfl).1.elim⟩ 123 | have ⟨_, he₁⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h3 c2 <| h1 _ this H4 124 | have ee₁ := h3.uniq henv (.refl henv hΔ) <| he₁.weakFV henv W hΔ 125 | have h4 := ee₁.symm.trans henv hΔ.toCtx h4 |>.trans henv hΔ.toCtx ee 126 | have h4 := (IsDefEqU.weakN_iff henv hΔ.toCtx W.toCtx).1 h4 127 | refine ⟨_, fun P hP he' => ?_, he, _, he₁, h4⟩ 128 | exact h1 _ 129 | ⟨(IsFVarUpSet.and_fvars hΔ.1.fvwf).1 hP, fun h => (hΔ.2.1 _ _ rfl).1.elim h.2⟩ 130 | (he'.mp (fun _ => .intro) he.fvarsIn) |>.mono fun _ => (·.1) 131 | 132 | theorem ConditionallyWHNF.fresh 133 | (henv : env.WF) 134 | (hΔ : VLCtx.WF env Us.length ((some (⟨ngen.curr⟩, deps), d) :: Δ)) 135 | (H : ConditionallyWHNF ngen env Us Δ e e₁) : 136 | ConditionallyWHNF ngen env Us ((some (⟨ngen.curr⟩, deps), d) :: Δ) e e₁ := by 137 | refine have ⟨c1, f1, c2, f2, H⟩ := H; ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 138 | have ⟨_, h1, h2, h3⟩ := H (H4.mp ?_ f1) 139 | · have W : VLCtx.FVLift Δ ((some (⟨ngen.curr⟩, deps), d) :: Δ) 0 (0 + d.depth) 0 := 140 | .skip_fvar _ _ .refl 141 | exact ⟨_, fun P hP => h1 _ hP.1, h2.weakFV henv W hΔ, h3.weakFV henv W hΔ⟩ 142 | · intro _ h1 h2; simp at h1; rcases h1 with rfl | h1 143 | · cases Nat.lt_irrefl _ (h2 _ rfl) 144 | · exact h1 145 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Typing/Expr.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Basic 2 | import Lean4Lean.Verify.NameGenerator 3 | import Lean4Lean.Verify.VLCtx 4 | import Lean4Lean.Verify.Axioms 5 | 6 | namespace Lean4Lean 7 | open Lean 8 | 9 | def Closed : Expr → (k :_:= 0) → Prop 10 | | .bvar i, k => i < k 11 | | .fvar _, _ | .sort .., _ | .const .., _ | .lit .., _ => True 12 | | .app f a, k => Closed f k ∧ Closed a k 13 | | .lam _ d b _, k 14 | | .forallE _ d b _, k => Closed d k ∧ Closed b (k+1) 15 | | .letE _ d v b _, k => Closed d k ∧ Closed v k ∧ Closed b (k+1) 16 | | .proj _ _ e, k | .mdata _ e, k => Closed e k 17 | | .mvar .., _ => False 18 | 19 | nonrec abbrev _root_.Lean.Expr.Closed := @Closed 20 | 21 | /-- This is very inefficient, only use for spec purposes -/ 22 | def _root_.Lean.Expr.fvarsList : Expr → List FVarId 23 | | .bvar _ | .sort .. | .const .. | .lit .. | .mvar .. => [] 24 | | .fvar fv => [fv] 25 | | .app f a => f.fvarsList ++ a.fvarsList 26 | | .lam _ d b _ 27 | | .forallE _ d b _ => d.fvarsList ++ b.fvarsList 28 | | .letE _ d v b _ => d.fvarsList ++ v.fvarsList ++ b.fvarsList 29 | | .proj _ _ e | .mdata _ e => e.fvarsList 30 | 31 | variable (fvars : FVarId → Prop) in 32 | def FVarsIn : Expr → Prop 33 | | .bvar _ => True 34 | | .fvar fv => fvars fv 35 | | .sort u => u.hasMVar' = false 36 | | .const _ us => ∀ u ∈ us, u.hasMVar' = false 37 | | .lit .. => True 38 | | .app f a => FVarsIn f ∧ FVarsIn a 39 | | .lam _ d b _ 40 | | .forallE _ d b _ => FVarsIn d ∧ FVarsIn b 41 | | .letE _ d v b _ => FVarsIn d ∧ FVarsIn v ∧ FVarsIn b 42 | | .proj _ _ e | .mdata _ e => FVarsIn e 43 | | .mvar .. => False 44 | 45 | nonrec abbrev _root_.Lean.Expr.FVarsIn := @FVarsIn 46 | 47 | def VLocalDecl.WF (env : VEnv) (U : Nat) (Γ : List VExpr) : VLocalDecl → Prop 48 | | .vlam type => env.IsType U Γ type 49 | | .vlet type value => env.HasType U Γ value type 50 | 51 | def VLCtx.FVWF : VLCtx → Prop 52 | | [] => True 53 | | (ofv, _) :: (Δ : VLCtx) => 54 | VLCtx.FVWF Δ ∧ (∀ fv deps, ofv = some (fv, deps) → fv ∉ Δ.fvars ∧ deps ⊆ Δ.fvars) 55 | 56 | variable (env : VEnv) (U : Nat) in 57 | def VLCtx.WF : VLCtx → Prop 58 | | [] => True 59 | | (ofv, d) :: (Δ : VLCtx) => 60 | VLCtx.WF Δ ∧ (∀ fv deps, ofv = some (fv, deps) → fv ∉ Δ.fvars ∧ deps ⊆ Δ.fvars) ∧ 61 | VLocalDecl.WF env U Δ.toCtx d 62 | 63 | def VLCtx.WF.fvwf : ∀ {Δ}, VLCtx.WF env U Δ → Δ.FVWF 64 | | [], h => h 65 | | _ :: _, ⟨h1, h2, _⟩ => ⟨h1.fvwf, h2⟩ 66 | 67 | def TrProj : ∀ (Γ : List VExpr) (structName : Name) (idx : Nat) (e : VExpr), VExpr → Prop := sorry 68 | 69 | variable (env : VEnv) (Us : List Name) in 70 | inductive TrExprS : VLCtx → Expr → VExpr → Prop 71 | | bvar : Δ.find? (.inl i) = some (e, A) → TrExprS Δ (.bvar i) e 72 | | fvar : Δ.find? (.inr fv) = some (e, A) → TrExprS Δ (.fvar fv) e 73 | | sort : VLevel.ofLevel Us u = some u' → TrExprS Δ (.sort u) (.sort u') 74 | | const : 75 | env.constants c = some (some ci) → 76 | us.mapM (VLevel.ofLevel Us) = some us' → 77 | us.length = ci.uvars → 78 | TrExprS Δ (.const c us) (.const c us') 79 | | app : 80 | env.HasType Us.length Δ.toCtx f' (.forallE A B) → 81 | env.HasType Us.length Δ.toCtx a' A → 82 | TrExprS Δ f f' → TrExprS Δ a a' → TrExprS Δ (.app f a) (.app f' a') 83 | | lam : 84 | env.IsType Us.length Δ.toCtx ty' → 85 | TrExprS Δ ty ty' → TrExprS ((none, .vlam ty') :: Δ) body body' → 86 | TrExprS Δ (.lam name ty body bi) (.lam ty' body') 87 | | forallE : 88 | env.IsType Us.length Δ.toCtx ty' → 89 | env.IsType Us.length (ty' :: Δ.toCtx) body' → 90 | TrExprS Δ ty ty' → TrExprS ((none, .vlam ty') :: Δ) body body' → 91 | TrExprS Δ (.forallE name ty body bi) (.forallE ty' body') 92 | | letE : 93 | env.HasType Us.length Δ.toCtx val' ty' → 94 | TrExprS Δ ty ty' → TrExprS Δ val val' → 95 | TrExprS ((none, .vlet ty' val') :: Δ) body body' → 96 | TrExprS Δ (.letE name ty val body nd) body' 97 | | lit : TrExprS Δ l.toConstructor e → TrExprS Δ (.lit l) e 98 | | mdata : TrExprS Δ e e' → TrExprS Δ (.mdata d e) e' 99 | | proj : TrExprS Δ e e' → TrProj Δ.toCtx s i e' e'' → TrExprS Δ (.proj s i e) e'' 100 | 101 | def TrExpr (env : VEnv) (Us : List Name) (Δ : VLCtx) (e : Expr) (e' : VExpr) : Prop := 102 | ∃ e₂, TrExprS env Us Δ e e₂ ∧ env.IsDefEqU Us.length Δ.toCtx e₂ e' 103 | 104 | def VExpr.bool : VExpr := .const ``Bool [] 105 | def VExpr.boolTrue : VExpr := .const ``Bool.true [] 106 | def VExpr.boolFalse : VExpr := .const ``Bool.false [] 107 | def VExpr.boolLit : Bool → VExpr 108 | | .false => .boolFalse 109 | | .true => .boolTrue 110 | 111 | def VExpr.nat : VExpr := .const ``Nat [] 112 | def VExpr.natZero : VExpr := .const ``Nat.zero [] 113 | def VExpr.natSucc : VExpr := .const ``Nat.succ [] 114 | def VExpr.natLit : Nat → VExpr 115 | | 0 => .natZero 116 | | n+1 => .app .natSucc (.natLit n) 117 | 118 | def VExpr.char : VExpr := .const ``Char [] 119 | def VExpr.string : VExpr := .const ``String [] 120 | def VExpr.stringMk : VExpr := .const ``String.mk [] 121 | def VExpr.listChar : VExpr := .app (.const ``List [.zero]) .char 122 | def VExpr.listCharNil : VExpr := .app (.const ``List.nil [.zero]) .char 123 | def VExpr.listCharCons : VExpr := .app (.const ``List.cons [.zero]) .char 124 | def VExpr.charOfNat : VExpr := .const ``Char.ofNat [] 125 | def VExpr.listCharLit : List Char → VExpr 126 | | [] => .listCharNil 127 | | a :: as => .app (.app .listCharCons (.app .charOfNat (.natLit a.toNat))) (.listCharLit as) 128 | 129 | def VExpr.trLiteral : Literal → VExpr 130 | | .natVal n => .natLit n 131 | | .strVal s => .app .stringMk (.listCharLit s.data) 132 | 133 | def VEnv.ReflectsNatNatNat (env : VEnv) (fc : Name) (f : Nat → Nat → Nat) := 134 | env.contains fc → 135 | ∀ a b, env.IsDefEqU 0 [] (.app (.app (.const fc []) (.natLit a)) (.natLit b)) (.natLit (f a b)) 136 | 137 | def VEnv.ReflectsNatNatBool (env : VEnv) (fc : Name) (f : Nat → Nat → Bool) := 138 | env.contains fc → 139 | ∀ a b, env.IsDefEqU 0 [] (.app (.app (.const fc []) (.natLit a)) (.natLit b)) (.boolLit (f a b)) 140 | 141 | structure VEnv.HasPrimitives (env : VEnv) : Prop where 142 | bool : env.contains ``Bool → env.contains ``Bool.false ∧ env.contains ``Bool.true 143 | boolFalse : env.constants ``Bool.false = some (some ci) → ci = { uvars := 0, type := .bool } 144 | boolTrue : env.constants ``Bool.true = some (some ci) → ci = { uvars := 0, type := .bool } 145 | nat : env.contains ``Nat → env.contains ``Nat.zero ∧ env.contains ``Nat.succ 146 | natZero : env.constants ``Nat.zero = some (some ci) → ci = { uvars := 0, type := .nat } 147 | natSucc : env.constants ``Nat.succ = some (some ci) → 148 | ci = { uvars := 0, type := .forallE .nat .nat } 149 | natAdd : env.ReflectsNatNatNat ``Nat.add Nat.add 150 | natSub : env.ReflectsNatNatNat ``Nat.sub Nat.sub 151 | natMul : env.ReflectsNatNatNat ``Nat.mul Nat.mul 152 | natPow : env.ReflectsNatNatNat ``Nat.pow Nat.pow 153 | natGcd : env.ReflectsNatNatNat ``Nat.gcd Nat.gcd 154 | natMod : env.ReflectsNatNatNat ``Nat.mod Nat.mod 155 | natDiv : env.ReflectsNatNatNat ``Nat.div Nat.div 156 | natBEq : env.ReflectsNatNatBool ``Nat.beq Nat.beq 157 | natBLE : env.ReflectsNatNatBool ``Nat.ble Nat.ble 158 | natLAnd : env.ReflectsNatNatNat ``Nat.land Nat.land 159 | natLOr : env.ReflectsNatNatNat ``Nat.lor Nat.lor 160 | natXor : env.ReflectsNatNatNat ``Nat.xor Nat.xor 161 | natShiftLeft : env.ReflectsNatNatNat ``Nat.shiftLeft Nat.shiftLeft 162 | natShiftRight : env.ReflectsNatNatNat ``Nat.shiftRight Nat.shiftRight 163 | string : env.contains ``String → env.contains ``String.mk ∧ 164 | env.HasType 0 [] .listCharNil .listChar ∧ 165 | env.HasType 0 [] .listCharCons (.forallE .char <| .forallE .listChar .listChar) ∧ 166 | env.HasType 0 [] .charOfNat (.forallE .nat .char) 167 | stringMk : env.constants ``String.mk = some (some ci) → 168 | ci = { uvars := 0, type := .forallE .listChar .string } 169 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/TypeChecker/Reduce.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.TypeChecker.Basic 2 | 3 | namespace Lean4Lean.TypeChecker.Inner 4 | open Lean hiding Environment Exception 5 | open Kernel 6 | 7 | theorem reduceNative.WF : 8 | (reduceNative env e).WF fun oe => ∀ e₁, oe = some e₁ → False := by 9 | unfold reduceNative; split <;> [skip; exact .pure nofun] 10 | split <;> [exact .throw; skip]; split <;> [exact .throw; exact .pure nofun] 11 | 12 | theorem rawNatLitExt?.WF {c : VContext} (H : rawNatLitExt? e = some n) (he : c.TrExprS e e') : 13 | c.venv.contains ``Nat ∧ e' = .natLit n := by 14 | have : c.TrExprS (.lit (.natVal n)) e' := by 15 | unfold rawNatLitExt? at H; split at H <;> rename_i h 16 | · cases H; exact .lit (he.eqv h) 17 | · unfold Expr.rawNatLit? at H; split at H <;> cases H; exact he 18 | have hn := this.lit_has_type c.Ewf c.hasPrimitives 19 | exact ⟨hn, this.unique (by trivial) (TrExprS.natLit c.hasPrimitives hn n).1⟩ 20 | 21 | def reduceBinNatOpG (guard : Nat → Nat → Prop) [DecidableRel guard] 22 | (f : Nat → Nat → Nat) (a b : Expr) : RecM (Option Expr) := do 23 | let some v1 := rawNatLitExt? (← whnf a) | return none 24 | let some v2 := rawNatLitExt? (← whnf b) | return none 25 | if guard v1 v2 then return none 26 | return some <| .lit <| .natVal <| f v1 v2 27 | 28 | theorem reduceBinNatOpG.WF {guard} [DecidableRel guard] {c : VContext} 29 | (he : c.TrExprS (.app (.app (.const fc ls) a) b) e') 30 | (hprim : Environment.primitives.contains fc) 31 | (heval : c.venv.ReflectsNatNatNat fc f) : 32 | RecM.WF c s (reduceBinNatOpG guard f a b) fun oe _ => ∀ e₁, oe = some e₁ → 33 | c.FVarsBelow (.app (.app (.const fc ls) a) b) e₁ ∧ c.TrExpr e₁ e' := by 34 | let .app hb1 hb2 hf hb := he 35 | let .app ha1 ha2 hf ha := hf 36 | let .const h1 h2 h3 := hf 37 | unfold reduceBinNatOpG 38 | refine (whnf.WF ha).bind fun a₁ _ _ ⟨a1, _, a2, a3⟩ => ?_ 39 | split <;> [rename_i v1 h; exact .pure nofun] 40 | obtain ⟨hn, rfl⟩ := rawNatLitExt?.WF h a2 41 | refine (whnf.WF hb).bind fun b₁ _ _ ⟨b1, _, b2, b3⟩ => ?_ 42 | split <;> [rename_i v2 h; exact .pure nofun] 43 | cases (rawNatLitExt?.WF h b2).2 44 | split <;> [exact .pure nofun; rename_i h] 45 | refine .pure ?_; rintro _ ⟨⟩; refine ⟨fun _ _ _ => trivial, ?_⟩ 46 | have ⟨ci, c1⟩ := c.trenv.find?_iff.2 ⟨_, h1⟩ 47 | have ⟨c2, c3⟩ := c.safePrimitives c1 hprim 48 | have ⟨_, d1, d2, d3⟩ := c.trenv.find?_uniq c1 h1 49 | simp [c3] at d2; simp [← d2] at h3; simp [h3] at h2; subst h2 50 | refine ⟨_, (TrExprS.natLit c.hasPrimitives hn _).1, ?_⟩ 51 | have := heval ⟨_, h1⟩ v1 v2 |>.instL (U' := c.lparams.length) (ls := []) nofun 52 | simp [VExpr.instL] at this 53 | refine this.weak0 c.Ewf (Γ := c.vlctx.toCtx) |>.symm.trans c.Ewf c.Δwf ?_ 54 | have a3 := a3.of_r c.Ewf c.Δwf ha2 55 | have b3 := b3.of_r c.Ewf c.Δwf hb2 56 | have := ha1.appDF a3 |>.toU.of_r c.Ewf c.Δwf hb1 57 | exact ⟨_, .appDF this b3⟩ 58 | 59 | theorem reduceBinNatPred.WF {c : VContext} 60 | (he : c.TrExprS (.app (.app (.const fc ls) a) b) e') 61 | (hprim : Environment.primitives.contains fc) 62 | (heval : c.venv.ReflectsNatNatBool fc f) : 63 | RecM.WF c s (reduceBinNatPred f a b) fun oe _ => ∀ e₁, oe = some e₁ → 64 | c.FVarsBelow (.app (.app (.const fc ls) a) b) e₁ ∧ c.TrExpr e₁ e' := by 65 | let .app hb1 hb2 hf hb := he 66 | let .app ha1 ha2 hf ha := hf 67 | let .const h1 h2 h3 := hf 68 | unfold reduceBinNatPred 69 | refine (whnf.WF ha).bind fun a₁ _ _ ⟨a1, _, a2, a3⟩ => ?_ 70 | split <;> [rename_i v1 h; exact .pure nofun]; cases (rawNatLitExt?.WF h a2).2 71 | refine (whnf.WF hb).bind fun b₁ _ _ ⟨b1, _, b2, b3⟩ => ?_ 72 | split <;> [rename_i v2 h; exact .pure nofun]; cases (rawNatLitExt?.WF h b2).2 73 | refine .pure ?_; rintro _ ⟨⟩; refine ⟨fun _ _ _ => .boolLit, ?_⟩ 74 | have ⟨ci, c1⟩ := c.trenv.find?_iff.2 ⟨_, h1⟩ 75 | have ⟨c2, c3⟩ := c.safePrimitives c1 hprim 76 | have ⟨_, d1, d2, d3⟩ := c.trenv.find?_uniq c1 h1 77 | simp [c3] at d2; simp [← d2] at h3; simp [h3] at h2; subst h2 78 | have := heval ⟨_, h1⟩ v1 v2 |>.instL (U' := c.lparams.length) (ls := []) nofun 79 | simp [VExpr.instL] at this 80 | refine ⟨_, (TrExprS.boolLit c.hasPrimitives ?_ _).1, ?_⟩ 81 | · let ⟨_, H⟩ := this 82 | exact VExpr.WF.boolLit_has_type c.Ewf c.hasPrimitives (Γ := []) trivial ⟨_, H.hasType.2⟩ 83 | refine this.weak0 c.Ewf (Γ := c.vlctx.toCtx) |>.symm.trans c.Ewf c.Δwf ?_ 84 | have a3 := a3.of_r c.Ewf c.Δwf ha2 85 | have b3 := b3.of_r c.Ewf c.Δwf hb2 86 | have := ha1.appDF a3 |>.toU.of_r c.Ewf c.Δwf hb1 87 | exact ⟨_, .appDF this b3⟩ 88 | 89 | theorem reduceNat.WF {c : VContext} (he : c.TrExprS e e') : 90 | RecM.WF c s (reduceNat e) fun oe _ => ∀ e₁, oe = some e₁ → 91 | c.FVarsBelow e e₁ ∧ c.TrExpr e₁ e' := by 92 | generalize hP : (fun oe => _) = P 93 | refine let prims := _; have hprims : Environment.primitives = .ofList prims := rfl; ?_ 94 | replace hprims {a} : Environment.primitives.contains a ↔ a ∈ prims := by 95 | simp [hprims, NameSet.contains, NameSet.ofList] 96 | unfold reduceNat 97 | extract_lets nargs F1 fn F2 98 | split <;> [exact hP ▸ .pure nofun; refine .pureBind ?_] 99 | unfold F2; split <;> (split <;> [skip; exact hP ▸ .pure nofun]) 100 | · rename_i _ h1 h2 101 | simp [nargs, Expr.getAppNumArgs_eq] at h1; subst fn 102 | let .app f a := e; simp [Expr.appFn!, Expr.eqv_const] at h2 ⊢; subst h2 103 | let .app ha1 ha2 hf ha := he 104 | let .const h1 h2 h3 := hf 105 | refine (whnf.WF ha).bind fun a₁ _ _ ⟨a1, _, a2, a3⟩ => ?_ 106 | split <;> [rename_i n h; exact hP ▸ .pure nofun] 107 | obtain ⟨hn, rfl⟩ := rawNatLitExt?.WF h a2 108 | refine hP ▸ .pure ?_; rintro _ ⟨⟩; refine ⟨fun _ _ _ => trivial, ?_⟩ 109 | have ⟨ci, c1⟩ := c.trenv.find?_iff.2 ⟨_, h1⟩ 110 | have ⟨c2, c3⟩ := c.safePrimitives c1 <| hprims.2 (by simp [prims]) 111 | have ⟨d1, d2, d3⟩ := c.trenv.find?_uniq c1 h1; cases h2 112 | refine have ⟨p1, p2⟩ := TrExprS.natLit c.hasPrimitives hn _; ⟨_, p1, ?_⟩ 113 | refine p2.toU.symm.trans c.Ewf c.Δwf ?_ 114 | exact ⟨_, ha1.appDF <| a3.of_r c.Ewf c.Δwf ha2⟩ 115 | · split <;> [rename_i f ls a b _ _ h2; exact hP ▸ .pure nofun] 116 | have hfun guard {g fc G} [DecidableRel guard] (hprim : fc ∈ prims) 117 | (heval : c.venv.ReflectsNatNatNat fc g) (hG : RecM.WF c s G P) : 118 | RecM.WF c s (do if f == fc then {return ← reduceBinNatOpG guard g a b}; G) P := by 119 | split <;> [rename_i h; exact hG] 120 | simp at h ⊢; subst h 121 | exact hP ▸ reduceBinNatOpG.WF he (hprims.2 hprim) heval 122 | have hpred {g fc G} (hprim : fc ∈ prims) 123 | (heval : c.venv.ReflectsNatNatBool fc g) (hG : RecM.WF c s G P) : 124 | RecM.WF c s (do if f == fc then {return ← reduceBinNatPred g a b}; G) P := by 125 | split <;> [rename_i h; exact hG] 126 | simp at h ⊢; subst h 127 | exact hP ▸ reduceBinNatPred.WF he (hprims.2 hprim) heval 128 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natAdd 129 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natSub 130 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natMul 131 | apply hfun _ (by simp [prims]) c.hasPrimitives.natPow 132 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natGcd 133 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natMod 134 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natDiv 135 | apply hpred (by simp [prims]) c.hasPrimitives.natBEq 136 | apply hpred (by simp [prims]) c.hasPrimitives.natBLE 137 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natLAnd 138 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natLOr 139 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natXor 140 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natShiftLeft 141 | apply hfun (fun _ _ => False) (by simp [prims]) c.hasPrimitives.natShiftRight 142 | exact hP ▸ .pure nofun 143 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Environment/Lemmas.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Std.SMap 2 | import Lean4Lean.Verify.Environment.Basic 3 | 4 | namespace Lean4Lean 5 | open Lean hiding Environment Exception 6 | open Kernel 7 | 8 | theorem TrConstant.mono {env env' : VEnv} (henv : env ≤ env') 9 | (H : TrConstant safety env ci ci') : TrConstant safety env' ci ci' := 10 | ⟨H.1, H.2.1, H.2.2.mono henv⟩ 11 | 12 | theorem TrConstVal.mono {env env' : VEnv} (henv : env ≤ env') 13 | (H : TrConstVal safety env ci ci') : TrConstVal safety env' ci ci' := 14 | ⟨H.1.mono henv, H.2⟩ 15 | 16 | theorem TrDefVal.mono {env env' : VEnv} (henv : env ≤ env') 17 | (H : TrDefVal safety env ci ci') : TrDefVal safety env' ci ci' := 18 | ⟨H.1.mono henv, H.2.mono henv⟩ 19 | 20 | variable (safety : DefinitionSafety) (env : VEnv) in 21 | inductive Aligned1 : ConstantInfo → Option VConstant → Prop where 22 | | block : ¬isAsSafeAs safety ci → Aligned1 ci none 23 | | const : TrConstant safety env ci ci' → Aligned1 ci (some ci') 24 | 25 | variable (safety : DefinitionSafety) in 26 | inductive Aligned : ConstMap → VEnv → Prop where 27 | | empty : Aligned {} .empty 28 | | const : Aligned C venv → Aligned1 safety venv ci ci' → 29 | venv.addConst n ci' = some venv' → ci.name = n → Aligned (C.insert n ci) venv' 30 | | defeq : Aligned C venv → Aligned C (venv.addDefEq df) 31 | 32 | theorem Aligned.map_wf' (H : Aligned safety C venv) : 33 | C.WF ∧ ∀ n, (C.find? n).isSome = (venv.constants n).isSome := by 34 | induction H with 35 | | empty => exact ⟨.empty, by simp [SMap.find?, VEnv.empty]⟩ 36 | | @const _ _ _ _ n _ _ _ eq _ ih => 37 | simp [VEnv.addConst] at eq; split at eq <;> cases eq 38 | have := ih.2 n; simp_all [-ih] 39 | refine ⟨.insert ih.1 _ _ this, fun n' => ?_⟩ 40 | simp [ih.1.find?_insert]; split <;> simp [ih.2] 41 | | defeq _ ih => exact ih 42 | 43 | theorem Aligned.map_wf (H : Aligned safety C venv) : C.WF := H.map_wf'.1 44 | theorem Aligned.find?_isSome (H : Aligned safety C venv) : 45 | (C.find? n).isSome = (venv.constants n).isSome := H.map_wf'.2 _ 46 | 47 | theorem Aligned.addQuot1 {Q : Prop} 48 | (H1 : ∀ c env, Aligned safety c env → P c env → Q) 49 | (C env) (wf : Aligned safety C env) (H2 : AddQuot1 n k ci P C env) : Q := by 50 | let ⟨_, _, h1, h2, h3, h4⟩ := H2 51 | exact H1 _ _ (wf.const (.const ⟨by cases safety <;> rfl, h2.2⟩) h3 rfl) h4 52 | 53 | nonrec theorem Aligned.addQuot (H : AddQuot C₁ C₂ venv₁ venv₂) 54 | (wf : Aligned safety C₁ venv₁) : Aligned safety C₂ venv₂ := by 55 | dsimp [AddQuot] at H 56 | refine (addQuot1 <| addQuot1 <| addQuot1 <| addQuot1 ?_) _ _ wf H 57 | rintro _ _ h ⟨rfl, rfl⟩; exact h.defeq 58 | 59 | theorem Aligned.addInduct (H : AddInduct C₁ venv₁ decl C₂ venv₂) : 60 | Aligned safety C₁ env₁ → Aligned safety C₂ env₂ := 61 | nomatch H 62 | 63 | theorem TrEnv'.aligned (H : TrEnv' safety C Q venv) : Aligned safety C venv := by 64 | induction H with 65 | | empty => exact .empty 66 | | block hs h _ ih => exact ih.const (.block hs) h rfl 67 | | «axiom» h1 _ h _ ih => exact ih.const (.const h1) h rfl 68 | | «opaque» h1 _ h _ ih => exact ih.const (.const h1.1.1) h rfl 69 | | defn h1 _ h _ ih => exact (ih.const (.const h1.1.1) h rfl).defeq 70 | | quot _ h _ ih => exact ih.addQuot h 71 | | induct _ h _ ih => exact ih.addInduct h 72 | 73 | theorem TrEnv'.map_wf (H : TrEnv' safety C Q venv) : C.WF := H.aligned.map_wf 74 | 75 | theorem Aligned.find? (H : Aligned safety C venv) 76 | (h : C.find? name = some ci) (hs : isAsSafeAs safety ci) : 77 | ∃ ci', venv.constants name = some (some ci') ∧ TrConstant safety venv ci ci' := by 78 | have mono {env₁ env₂} (H : env₁.LE env₂) : 79 | (∃ ci', env₁.constants name = some (some ci') ∧ TrConstant safety env₁ ci ci') → 80 | (∃ ci', env₂.constants name = some (some ci') ∧ TrConstant safety env₂ ci ci') 81 | | ⟨_, h1, h2⟩ => ⟨_, H.constants h1, h2.mono H⟩ 82 | induction H with 83 | | empty => simp [SMap.find?] at h 84 | | const h1 h2 h3 h4 ih => 85 | have := VEnv.addConst_le h3 86 | rw [h1.map_wf.find?_insert] at h; split at h 87 | · cases h; rename_i n _ h; simp at h; subst n h 88 | simp [VEnv.addConst] at h3; split at h3 <;> cases h3 89 | cases h2; · contradiction 90 | simp; rename_i h'; exact h'.mono this 91 | · let ⟨_, h1, h2⟩ := ih h; exact ⟨_, this.constants h1, h2.mono this⟩ 92 | | defeq h1 ih => let ⟨_, h1, h2⟩ := ih h; exact ⟨_, h1, h2.mono VEnv.addDefEq_le⟩ 93 | 94 | theorem Aligned.find?_uniq (H : Aligned safety C venv) 95 | (h : C.find? name = some ci) (hs : venv.constants name = some (some ci')) : 96 | ci.name = name ∧ TrConstant safety venv ci ci' := by 97 | induction H with 98 | | empty => simp [SMap.find?] at h 99 | | const h1 h2 h3 h4 ih => 100 | have := VEnv.addConst_le h3 101 | simp [VEnv.addConst] at h3; split at h3 <;> cases h3 102 | simp [h1.map_wf.find?_insert] at h hs; revert h hs; split 103 | · rintro ⟨⟩ ⟨⟩; rename_i n _ _ _; subst n 104 | let .const h2 := h2; exact ⟨h4, h2.mono this⟩ 105 | · intro hs h; let ⟨h1, h2⟩ := ih h hs; exact ⟨h1, h2.mono this⟩ 106 | | defeq h1 ih => let ⟨h1, h2⟩ := ih h hs; exact ⟨h1, h2.mono VEnv.addDefEq_le⟩ 107 | 108 | theorem TrEnv.find?_iff (H : TrEnv safety env venv) : 109 | (∃ ci, env.find? name = some ci) ↔ ∃ oci, venv.constants name = some oci := by 110 | conv => enter [1,1,_,1]; apply H.map_wf.find?'_eq_find? 111 | erw [← Option.isSome_iff_exists, H.aligned.find?_isSome, Option.isSome_iff_exists] 112 | 113 | theorem TrEnv.find? (H : TrEnv safety env venv) 114 | (h : env.find? name = some ci) (hs : isAsSafeAs safety ci) : 115 | ∃ ci', venv.constants name = some (some ci') ∧ TrConstant safety venv ci ci' := 116 | H.aligned.find? (H.map_wf.find?'_eq_find? _ ▸ h) hs 117 | 118 | theorem TrEnv.find?_uniq (H : TrEnv safety env venv) 119 | (h : env.find? name = some ci) (hs : venv.constants name = some (some ci')) : 120 | ci.name = name ∧ TrConstant safety venv ci ci' := 121 | H.aligned.find?_uniq (H.map_wf.find?'_eq_find? _ ▸ h) hs 122 | 123 | theorem TrEnv'.of_value (H : TrEnv' safety C Q venv) (h : C.find? name = some ci) 124 | (hs : isAsSafeAs safety ci) (hv : ci.value? = some v) : 125 | TrExpr venv ci.levelParams [] v (.const ci.name (VLevel.params ci.levelParams.length)) := by 126 | have {C n ci'} (hC : C.WF) : 127 | (SMap.insert C n ci').find? name = some ci → 128 | C.find? name = some ci ∨ n = name ∧ ci' = ci := by 129 | rw [hC.find?_insert]; simp; split <;> simp +contextual [*] 130 | induction H with 131 | | empty => simp [SMap.find?] at h 132 | | block _ h1 H ih | «axiom» _ _ h1 H ih | «opaque» _ _ h1 H ih => 133 | obtain h | ⟨rfl, rfl⟩ := this H.map_wf h 134 | · exact (ih h).mono (VEnv.addConst_le h1) 135 | · contradiction 136 | | defn h2 h3 h1 H ih => 137 | have' le := (VEnv.addConst_le h1).trans VEnv.addDefEq_le 138 | obtain h | ⟨rfl, rfl⟩ := this H.map_wf h 139 | · exact (ih h).mono le 140 | · cases hv 141 | have := VEnv.IsDefEq.extra0 VEnv.addDefEq_self <| 142 | (H.defn h2 h3 h1).wf.ordered.defEqWF VEnv.addDefEq_self 143 | let ⟨⟨⟨b1, b2, b3⟩, b4⟩, b5⟩ := h2 144 | refine ⟨_, b5.mono le, b2.symm ▸ b4.symm ▸ ⟨_, this.symm⟩⟩ 145 | | quot _ h1 H ih => 146 | suffices ∀ {n k ci' P}, (∀ C env, Aligned safety C env → P C env → C.find? name = some ci) → 147 | ∀ C env, Aligned safety C env → AddQuot1 n k ci' P C env → C.find? name = some ci by 148 | refine (ih <| this (this <| this <| this ?_) _ _ H.aligned h1).mono h1.le 149 | rintro _ _ _ ⟨rfl, rfl⟩; exact h 150 | rintro n k ci' P ih C env wf ⟨_, h1, _, h2, h3, h4⟩ 151 | have wf' := wf.const (.const ⟨by cases safety <;> rfl, h2.2⟩) h3 rfl 152 | obtain h | ⟨rfl, rfl⟩ := this wf.map_wf (ih _ _ wf' h4) 153 | · exact h 154 | · contradiction 155 | | induct _ h1 H ih => cases h1 156 | 157 | nonrec theorem TrEnv.of_value (H : TrEnv safety env venv) (h : env.find? name = some ci) 158 | (hs : isAsSafeAs safety ci) (hv : ci.value? = some v) : 159 | TrExpr venv ci.levelParams [] v (.const ci.name (VLevel.params ci.levelParams.length)) := 160 | H.of_value (by rwa [← H.map_wf.find?'_eq_find?]) hs hv 161 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/UniqueTyping.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Lemmas 2 | import Lean4Lean.Theory.Typing.EnvLemmas 3 | 4 | /-! 5 | A bunch of important structural theorems which we can't prove :( 6 | -/ 7 | 8 | namespace Lean4Lean 9 | namespace VEnv 10 | 11 | theorem IsDefEq.uniq (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 12 | (h1 : env.IsDefEq U Γ e₁ e₂ A) (h2 : env.IsDefEq U Γ e₂ e₃ B) : 13 | ∃ u, env.IsDefEq U Γ A B (.sort u) := sorry 14 | 15 | theorem IsDefEq.uniqU (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 16 | (h1 : env.IsDefEq U Γ e₁ e₂ A) (h2 : env.IsDefEq U Γ e₂ e₃ B) : 17 | env.IsDefEqU U Γ A B := let ⟨_, h⟩ := h1.uniq henv hΓ h2; ⟨_, h⟩ 18 | 19 | variable! (henv : VEnv.WF env) in 20 | theorem IsDefEqU.weakN (W : Ctx.LiftN n k Γ Γ') (H : env.IsDefEqU U Γ e1 e2) : 21 | env.IsDefEqU U Γ' (e1.liftN n k) (e2.liftN n k) := let ⟨_, H⟩ := H; ⟨_, H.weakN henv W⟩ 22 | 23 | variable! (henv : VEnv.WF env) in 24 | theorem IsDefEqU.weak' (W : Ctx.Lift' n Γ Γ') (H : env.IsDefEqU U Γ e1 e2) : 25 | env.IsDefEqU U Γ' (e1.lift' n) (e2.lift' n) := let ⟨_, H⟩ := H; ⟨_, H.weak' henv W⟩ 26 | 27 | theorem isDefEq_iff (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) : 28 | env.IsDefEq U Γ e₁ e₂ A ↔ 29 | env.HasType U Γ e₁ A ∧ env.HasType U Γ e₂ A ∧ env.IsDefEqU U Γ e₁ e₂ := by 30 | refine ⟨fun h => ⟨h.hasType.1, h.hasType.2, _, h⟩, fun ⟨_, h2, _, h3⟩ => ?_⟩ 31 | have ⟨_, h⟩ := h3.uniq henv hΓ h2 32 | exact h.defeqDF h3 33 | 34 | theorem IsDefEq.trans_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 35 | (h₁ : env.IsDefEq U Γ e₁ e₂ A) (h₂ : env.IsDefEq U Γ e₂ e₃ B) : 36 | env.IsDefEq U Γ e₁ e₃ B := by 37 | have ⟨_, h⟩ := h₁.uniq henv hΓ h₂ 38 | exact .trans (.defeqDF h h₁) h₂ 39 | 40 | theorem IsDefEq.trans_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 41 | (h₁ : env.IsDefEq U Γ e₁ e₂ A) (h₂ : env.IsDefEq U Γ e₂ e₃ B) : 42 | env.IsDefEq U Γ e₁ e₃ A := by 43 | have ⟨_, h⟩ := h₁.uniq henv hΓ h₂ 44 | exact h₁.trans (.defeqDF (.symm h) h₂) 45 | 46 | theorem IsDefEqU.defeqDF (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 47 | (h₁ : env.IsDefEqU U Γ A B) (h₂ : env.IsDefEq U Γ e₁ e₂ A) : 48 | env.IsDefEq U Γ e₁ e₂ B := by 49 | have ⟨_, h₁⟩ := h₁ 50 | have ⟨_, hA⟩ := h₂.isType henv hΓ 51 | exact .defeqDF (hA.trans_l henv hΓ h₁) h₂ 52 | 53 | theorem IsDefEqU.of_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 54 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.HasType U Γ e₁ A) : 55 | env.IsDefEq U Γ e₁ e₂ A := let ⟨_, h⟩ := h1; h2.trans_l henv hΓ h 56 | 57 | theorem HasType.defeqU_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 58 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.HasType U Γ e₁ A) : 59 | env.HasType U Γ e₂ A := (h1.of_l henv hΓ h2).hasType.2 60 | 61 | theorem IsType.defeqU_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 62 | (h1 : env.IsDefEqU U Γ A₁ A₂) (h2 : env.IsType U Γ A₁) : 63 | env.IsType U Γ A₂ := h2.imp fun _ h2 => h2.defeqU_l henv hΓ h1 64 | 65 | theorem IsDefEqU.of_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 66 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.HasType U Γ e₂ A) : 67 | env.IsDefEq U Γ e₁ e₂ A := (h1.symm.of_l henv hΓ h2).symm 68 | 69 | theorem HasType.defeqU_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 70 | (h1 : env.IsDefEqU U Γ A₁ A₂) (h2 : env.HasType U Γ e A₁) : 71 | env.HasType U Γ e A₂ := h1.defeqDF henv hΓ h2 72 | 73 | theorem IsDefEqU.trans (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 74 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.IsDefEqU U Γ e₂ e₃) : 75 | env.IsDefEqU U Γ e₁ e₃ := h1.imp fun _ h1 => let ⟨_, h2⟩ := h2; h1.trans_l henv hΓ h2 76 | 77 | theorem IsDefEqU.sort_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 78 | (h1 : env.IsDefEqU U Γ (.sort u) (.sort v)) : u ≈ v := sorry 79 | 80 | theorem IsDefEqU.forallE_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 81 | (h1 : env.IsDefEqU U Γ (.forallE A B) (.forallE A' B')) : 82 | (∃ u, env.IsDefEq U Γ A A' (.sort u)) ∧ ∃ u, env.IsDefEq U (A::Γ) B B' (.sort u) := sorry 83 | 84 | theorem IsDefEqU.sort_forallE_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) : 85 | ¬env.IsDefEqU U Γ (.sort u) (.forallE A B) := sorry 86 | 87 | variable! (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) in 88 | theorem IsDefEqU.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 89 | env.IsDefEqU U Γ' (e1.liftN n k) (e2.liftN n k) ↔ env.IsDefEqU U Γ e1 e2 := by 90 | refine ⟨fun h => have := henv; have := hΓ; sorry, fun h => h.weakN henv W⟩ 91 | 92 | variable! (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) in 93 | theorem _root_.Lean4Lean.VExpr.WF.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 94 | VExpr.WF env U Γ' (e.liftN n k) ↔ VExpr.WF env U Γ e := IsDefEqU.weakN_iff henv hΓ W 95 | 96 | theorem IsDefEq.skips (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) 97 | (W : Ctx.LiftN n k Γ Γ') 98 | (H : env.IsDefEq U Γ' e₁ e₂ A) (h1 : e₁.Skips n k) (h2 : e₂.Skips n k) : 99 | ∃ B, env.IsDefEq U Γ' e₁ e₂ B ∧ B.Skips n k := by 100 | obtain ⟨e₁, rfl⟩ := VExpr.skips_iff_exists.1 h1 101 | obtain ⟨e₂, rfl⟩ := VExpr.skips_iff_exists.1 h2 102 | have ⟨_, H⟩ := (IsDefEqU.weakN_iff henv hΓ W).1 ⟨_, H⟩ 103 | exact ⟨_, H.weakN henv W, .liftN⟩ 104 | 105 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) (hΓ : OnCtx Γ (env.IsType U)) in 106 | theorem IsDefEq.weakN_iff' (W : Ctx.LiftN n k Γ Γ') : 107 | env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k) ↔ env.IsDefEq U Γ e1 e2 A := by 108 | refine ⟨fun h => ?_, fun h => h.weakN henv W⟩ 109 | have ⟨_, H⟩ := (IsDefEqU.weakN_iff henv hΓ' W).1 ⟨_, h⟩ 110 | refine IsDefEqU.defeqDF henv hΓ ?_ H 111 | exact (IsDefEqU.weakN_iff henv hΓ' W).1 <| (H.weakN henv W).uniqU henv hΓ' h.symm 112 | 113 | variable! (henv : VEnv.WF env) in 114 | theorem _root_.Lean4Lean.OnCtx.weakN_inv 115 | (W : Ctx.LiftN n k Γ Γ') (H : OnCtx Γ' (env.IsType U)) : OnCtx Γ (env.IsType U) := by 116 | induction W with 117 | | zero As h => 118 | clear h 119 | induction As with 120 | | nil => exact H 121 | | cons A As ih => exact ih H.1 122 | | succ W ih => 123 | let ⟨H1, _, H2⟩ := H 124 | exact ⟨ih H1, _, (IsDefEq.weakN_iff' henv H1 (ih H1) W).1 H2⟩ 125 | 126 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 127 | theorem IsDefEq.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 128 | env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k) ↔ env.IsDefEq U Γ e1 e2 A := 129 | IsDefEq.weakN_iff' henv hΓ' (hΓ'.weakN_inv henv W) W 130 | 131 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 132 | theorem HasType.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 133 | env.HasType U Γ' (e.liftN n k) (A.liftN n k) ↔ env.HasType U Γ e A := 134 | IsDefEq.weakN_iff henv hΓ' W 135 | 136 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 137 | theorem IsType.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 138 | env.IsType U Γ' (A.liftN n k) ↔ env.IsType U Γ A := 139 | exists_congr fun _ => HasType.weakN_iff henv hΓ' W (A := .sort _) 140 | 141 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 142 | theorem HasType.skips (W : Ctx.LiftN n k Γ Γ') 143 | (h1 : env.HasType U Γ' e A) (h2 : e.Skips n k) : ∃ B, env.HasType U Γ' e B ∧ B.Skips n k := 144 | IsDefEq.skips henv hΓ' W h1 h2 h2 145 | 146 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 147 | theorem IsDefEqU.weak'_iff (W : Ctx.Lift' l Γ Γ') : 148 | env.IsDefEqU U Γ' (e1.lift' l) (e2.lift' l) ↔ env.IsDefEqU U Γ e1 e2 := by 149 | generalize e : l.depth = n 150 | induction n generalizing l Γ' with 151 | | zero => simp [VExpr.lift'_depth_zero e, W.depth_zero e] 152 | | succ n ih => 153 | obtain ⟨l, k, rfl, rfl⟩ := Lift.depth_succ e 154 | have ⟨Γ₁, W1, W2⟩ := W.of_cons_skip 155 | rw [Lift.consN_skip_eq, VExpr.lift'_comp, VExpr.lift'_comp, 156 | ← Lift.skipN_one, VExpr.lift'_consN_skipN, VExpr.lift'_consN_skipN, 157 | weakN_iff henv hΓ' W2, ih (hΓ'.weakN_inv henv W2) W1 Lift.depth_consN] 158 | 159 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 160 | theorem IsDefEq.weak'_iff (W : Ctx.Lift' l Γ Γ') : 161 | env.IsDefEq U Γ' (e1.lift' l) (e2.lift' l) (A.lift' l) ↔ env.IsDefEq U Γ e1 e2 A := by 162 | generalize e : l.depth = n 163 | induction n generalizing l Γ' with 164 | | zero => simp [VExpr.lift'_depth_zero e, W.depth_zero e] 165 | | succ n ih => 166 | obtain ⟨l, k, rfl, rfl⟩ := Lift.depth_succ e 167 | have ⟨Γ₁, W1, W2⟩ := W.of_cons_skip 168 | rw [Lift.consN_skip_eq, VExpr.lift'_comp, VExpr.lift'_comp, VExpr.lift'_comp, 169 | ← Lift.skipN_one, VExpr.lift'_consN_skipN, VExpr.lift'_consN_skipN, VExpr.lift'_consN_skipN, 170 | weakN_iff henv hΓ' W2, ih (hΓ'.weakN_inv henv W2) W1 Lift.depth_consN] 171 | 172 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 173 | theorem HasType.weak'_iff (W : Ctx.Lift' l Γ Γ') : 174 | env.HasType U Γ' (e.lift' l) (A.lift' l) ↔ env.HasType U Γ e A := 175 | IsDefEq.weak'_iff henv hΓ' W 176 | 177 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 178 | theorem IsType.weak'_iff (W : Ctx.Lift' l Γ Γ') : 179 | env.IsType U Γ' (e.lift' l) ↔ env.IsType U Γ e := 180 | exists_congr fun _ => HasType.weak'_iff henv hΓ' W (A := .sort _) 181 | 182 | variable! (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) in 183 | theorem _root_.Lean4Lean.VExpr.WF.weak'_iff (W : Ctx.Lift' l Γ Γ') : 184 | VExpr.WF env U Γ' (e.lift' l) ↔ VExpr.WF env U Γ e := IsDefEqU.weak'_iff henv hΓ W 185 | -------------------------------------------------------------------------------- /Lean4Lean/Experimental/ParallelReduction.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Strong 2 | import Lean4Lean.Theory.Typing.NormalEq 3 | 4 | inductive ReflTransGen (R : α → α → Prop) (a : α) : α → Prop where 5 | | zero : ReflTransGen R a a 6 | | trans : ReflTransGen R a b → R b c → ReflTransGen R a c 7 | 8 | namespace Lean4Lean 9 | 10 | open VExpr 11 | 12 | section 13 | set_option hygiene false 14 | variable (TY : Typing) 15 | local notation:65 Γ " ⊢ " e1 " ≫ " e2:30 => ParRed Γ e1 e2 16 | local notation:65 Γ " ⊢ " e1 " ⋙ " e2:30 => CParRed Γ e1 e2 17 | 18 | inductive ParRed : List VExpr → VExpr → VExpr → Prop where 19 | | bvar : Γ ⊢ .bvar i ≫ .bvar i 20 | | sort : Γ ⊢ .sort u ≫ .sort u 21 | | const : Γ ⊢ .const c ls ≫ .const c ls 22 | | app : Γ ⊢ f ≫ f' → Γ ⊢ a ≫ a' → Γ ⊢ .app f a ≫ .app f' a' 23 | | lam : Γ ⊢ A ≫ A' → A::Γ ⊢ body ≫ body' → Γ ⊢ .lam A body ≫ .lam A' body' 24 | | forallE : Γ ⊢ A ≫ A' → A::Γ ⊢ B ≫ B' → Γ ⊢ .forallE A B ≫ .forallE A' B' 25 | | beta : A::Γ ⊢ e₁ ≫ e₁' → Γ ⊢ e₂ ≫ e₂' → Γ ⊢ .app (.lam A e₁) e₂ ≫ e₁'.inst e₂' 26 | | extra : TY.Pat p r → p.Matches e m1 m2 → r.2.OK (TY.IsDefEq Γ) m1 m2 → 27 | (∀ a, Γ ⊢ m2 a ≫ m2' a) → Γ ⊢ e ≫ r.1.apply m1 m2' 28 | 29 | def IsNeutral (Γ : List VExpr) (e : VExpr) : Prop := 30 | (∀ A e₁ e₂, e ≠ .app (.lam A e₁) e₂) ∧ 31 | (∀ p r m1 m2, TY.Pat p r → p.Matches e m1 m2 → ¬r.2.OK (TY.IsDefEq Γ) m1 m2) 32 | 33 | inductive CParRed : List VExpr → VExpr → VExpr → Prop where 34 | | bvar : Γ ⊢ .bvar i ⋙ .bvar i 35 | | sort : Γ ⊢ .sort u ⋙ .sort u 36 | | const : IsNeutral TY Γ (.const c ls) → Γ ⊢ .const c ls ⋙ .const c ls 37 | | app : IsNeutral TY Γ (.app f a) → Γ ⊢ f ⋙ f' → Γ ⊢ a ⋙ a' → Γ ⊢ .app f a ⋙ .app f' a' 38 | | lam : Γ ⊢ A ⋙ A' → A::Γ ⊢ body ⋙ body' → Γ ⊢ .lam A body ⋙ .lam A' body' 39 | | forallE : Γ ⊢ A ⋙ A' → A::Γ ⊢ B ⋙ B' → Γ ⊢ .forallE A B ⋙ .forallE A' B' 40 | | beta : A::Γ ⊢ e₁ ⋙ e₁' → Γ ⊢ e₂ ⋙ e₂' → Γ ⊢ .app (.lam A e₁) e₂ ⋙ e₁'.inst e₂' 41 | | extra : TY.Pat p r → p.Matches e m1 m2 → r.2.OK (TY.IsDefEq Γ) m1 m2 → 42 | (∀ a, Γ ⊢ m2 a ⋙ m2' a) → Γ ⊢ e ⋙ r.1.apply m1 m2' 43 | 44 | end 45 | 46 | variable {TY : Typing} 47 | 48 | protected theorem ParRed.rfl : ∀ {e}, ParRed TY Γ e e 49 | | .bvar .. => .bvar 50 | | .sort .. => .sort 51 | | .const .. => .const 52 | | .app .. => .app ParRed.rfl ParRed.rfl 53 | | .lam .. => .lam ParRed.rfl ParRed.rfl 54 | | .forallE .. => .forallE ParRed.rfl ParRed.rfl 55 | 56 | theorem ParRed.weakN (W : Ctx.LiftN n k Γ Γ') (H : ParRed TY Γ e1 e2) : 57 | ParRed TY Γ' (e1.liftN n k) (e2.liftN n k) := by 58 | induction H generalizing k Γ' with 59 | | bvar | sort | const => exact .rfl 60 | | app _ _ ih1 ih2 => exact .app (ih1 W) (ih2 W) 61 | | lam _ _ ih1 ih2 => exact .lam (ih1 W) (ih2 W.succ) 62 | | forallE _ _ ih1 ih2 => exact .forallE (ih1 W) (ih2 W.succ) 63 | | beta _ _ ih1 ih2 => 64 | simp [liftN, liftN_inst_hi] 65 | exact .beta (ih1 W.succ) (ih2 W) 66 | | extra h1 h2 h3 _ ih => 67 | rw [Pattern.RHS.liftN_apply] 68 | refine .extra h1 (Pattern.matches_liftN.2 ⟨_, h2, funext_iff.1 rfl⟩) 69 | (h3.weakN W) (fun a => ih _ W) 70 | 71 | variable! (H₀ : ParRed TY Γ₀ a1 a2) (H₀' : TY.HasType Γ₀ a1 A₀) in 72 | theorem ParRed.instN (W : Ctx.InstN Γ₀ a1 A₀ k Γ₁ Γ) 73 | (H : ParRed TY Γ₁ e1 e2) : ParRed TY Γ (e1.inst a1 k) (e2.inst a2 k) := by 74 | induction H generalizing Γ k with 75 | | @bvar _ i => 76 | dsimp [inst] 77 | induction W generalizing i with 78 | | zero => 79 | cases i with simp 80 | | zero => exact H₀ 81 | | succ h => exact .rfl 82 | | succ _ ih => 83 | cases i with simp 84 | | zero => exact .rfl 85 | | succ h => exact ih.weakN .one 86 | | sort | const => exact .rfl 87 | | app _ _ ih1 ih2 => exact .app (ih1 W) (ih2 W) 88 | | lam _ _ ih1 ih2 => exact .lam (ih1 W) (ih2 W.succ) 89 | | forallE _ _ ih1 ih2 => exact .forallE (ih1 W) (ih2 W.succ) 90 | | beta _ _ ih1 ih2 => 91 | simp [inst, inst0_inst_hi] 92 | exact .beta (ih1 W.succ) (ih2 W) 93 | | extra h1 h2 h3 _ ih => 94 | rw [Pattern.RHS.instN_apply] 95 | exact .extra h1 (Pattern.matches_instN h2) (h3.instN W H₀') (fun a => ih _ W) 96 | 97 | theorem ParRed.defeq (H : ParRed TY Γ e e') (he : TY.HasType Γ e A) : TY.IsDefEq Γ e e' := by 98 | induction H generalizing A with 99 | | bvar | sort | const => exact TY.refl he 100 | | app _ _ ih1 ih2 => 101 | have ⟨_, _, h1, h2⟩ := TY.app_inv he 102 | exact TY.appDF h1 (ih1 h1) h2 (ih2 h2) 103 | | lam _ _ ih1 ih2 => 104 | have ⟨_, _, h1, h2⟩ := TY.lam_inv he 105 | exact TY.lamDF h1 (ih1 h1) (ih2 h2) 106 | | forallE _ _ ih1 ih2 => 107 | have ⟨_, _, h1, h2⟩ := TY.forallE_inv he 108 | exact TY.forallEDF h1 (ih1 h1) h2 (ih2 h2) 109 | | beta _ _ ih1 ih2 => 110 | have ⟨_, _, hf, ha⟩ := TY.app_inv he 111 | have ⟨_, _, hA, hb⟩ := TY.lam_inv hf 112 | have hf' := TY.lam hA hb 113 | replace ha := TY.defeq_r (TY.forallE_defInv (TY.uniq hf hf')).1 ha 114 | exact TY.trans (TY.appDF hf' (TY.lamDF hA (TY.refl hA) (ih1 hb)) ha (ih2 ha)) 115 | (TY.beta (TY.defeq_l (ih1 hb) hb) (TY.defeq_l (ih2 ha) ha)) 116 | | @extra p r e m1 m2 Γ m2' h1 h2 h3 _ ih => 117 | exact TY.trans (TY.pat_wf h1 h2 he h3) <| 118 | .apply_pat ih (TY.defeq_l (TY.pat_wf h1 h2 he h3) he) 119 | 120 | theorem ParRed.hasType (H : ParRed TY Γ e e') (he : TY.HasType Γ e A) : TY.HasType Γ e' A := 121 | TY.defeq_l (H.defeq he) he 122 | 123 | theorem ParRed.defeqDFC (W : IsDefEqCtx TY.IsDefEq Γ₀ Γ₁ Γ₂) 124 | (h : TY.HasType Γ₁ e1 A) (H : ParRed TY Γ₁ e1 e2) : ParRed TY Γ₂ e1 e2 := by 125 | induction H generalizing Γ₂ A with 126 | | bvar => exact .bvar 127 | | sort => exact .sort 128 | | const => exact .const 129 | | app _ _ ih1 ih2 => 130 | have ⟨_, _, hf, ha⟩ := TY.app_inv h 131 | exact .app (ih1 W hf) (ih2 W ha) 132 | | lam _ _ ih1 ih2 => 133 | have ⟨_, _, hA, he⟩ := TY.lam_inv h 134 | exact .lam (ih1 W hA) (ih2 (W.succ (TY.refl hA)) he) 135 | | forallE _ _ ih1 ih2 => 136 | have ⟨_, _, hA, hB⟩ := TY.forallE_inv h 137 | exact .forallE (ih1 W hA) (ih2 (W.succ (TY.refl hA)) hB) 138 | | beta _ _ ih1 ih2 => 139 | have ⟨_, _, hf, ha⟩ := TY.app_inv h 140 | have ⟨_, _, hA, hb⟩ := TY.lam_inv hf 141 | exact .beta (ih1 (W.succ (TY.refl hA)) hb) (ih2 W ha) 142 | | @extra p r e m1 m2 Γ m2' h1 h2 h3 _ ih => 143 | exact .extra h1 h2 (h3.map fun a b h => TY.isDefEq_DFC W h) fun a => 144 | let ⟨_, h⟩ := h2.hasType h a; ih a W h 145 | 146 | 147 | theorem ParRed.apply_pat {p : Pattern} (r : p.RHS) {m1 m2 m3} 148 | (H : ∀ a, ParRed TY Γ (m2 a) (m3 a)) : ParRed TY Γ (r.apply m1 m2) (r.apply m1 m3) := by 149 | match r with 150 | | .fixed .. => exact .rfl 151 | | .app f a => exact .app (apply_pat f H) (apply_pat a H) 152 | | .var f => exact H _ 153 | 154 | theorem CParRed.toParRed (H : CParRed TY Γ e e') : ParRed TY Γ e e' := by 155 | induction H with 156 | | bvar => exact .bvar 157 | | sort => exact .sort 158 | | const => exact .const 159 | | app _ _ _ ih1 ih2 => exact .app ih1 ih2 160 | | lam _ _ ih1 ih2 => exact .lam ih1 ih2 161 | | forallE _ _ ih1 ih2 => exact .forallE ih1 ih2 162 | | beta _ _ ih1 ih2 => exact .beta ih1 ih2 163 | | extra h1 h2 h3 _ ih3 => exact .extra h1 h2 h3 ih3 164 | 165 | theorem CParRed.exists (H : TY.HasType Γ e A) : ∃ e', CParRed TY Γ e e' := by 166 | induction e using VExpr.brecOn generalizing Γ A with | _ e e_ih => ?_ 167 | have neut e 168 | (H' : TY.HasType Γ e A) 169 | (e_ih : VExpr.below (motive := fun {e} => 170 | ∀ {Γ A}, TY.HasType Γ e A → ∃ e', CParRed TY Γ e e') e) 171 | (hn : IsNeutral TY Γ e → ∃ e', CParRed TY Γ e e') : ∃ e', CParRed TY Γ e e' := by 172 | by_cases h : ∃ A e₁ e₂, e = .app (.lam A e₁) e₂ 173 | · obtain ⟨A, e, a, rfl⟩ := h 174 | have ⟨_, _, hf, ha⟩ := TY.app_inv H' 175 | have ⟨_, _, _, he⟩ := TY.lam_inv hf 176 | have ⟨_, he⟩ := e_ih.1.2.2.1 he 177 | have ⟨_, ha⟩ := e_ih.2.1 ha 178 | exact ⟨_, .beta he ha⟩ 179 | by_cases h' : ∃ p r m1 m2, TY.Pat p r ∧ p.Matches e m1 m2 ∧ r.2.OK (TY.IsDefEq Γ) m1 m2 180 | · let ⟨p, r, m1, m2, h1, hp2, hp3⟩ := h' 181 | suffices ∃ m3 : p.Path.2 → VExpr, ∀ a, CParRed TY Γ (m2 a) (m3 a) from 182 | let ⟨_, h3⟩ := this; ⟨_, .extra h1 hp2 hp3 h3⟩ 183 | clear H hn h h' r h1 hp3 184 | induction p generalizing e A with 185 | | const => exact ⟨nofun, nofun⟩ 186 | | app f a ih1 ih2 => 187 | let .app hm1 hm2 := hp2 188 | have ⟨_, _, H1, H2⟩ := TY.app_inv H' 189 | have ⟨m2l, hl⟩ := ih1 _ H1 e_ih.1.2 _ _ hm1 190 | have ⟨m2r, hr⟩ := ih2 _ H2 e_ih.2.2 _ _ hm2 191 | exact ⟨Sum.elim m2l m2r, Sum.rec hl hr⟩ 192 | | var _ ih => 193 | let .var hm1 := hp2 194 | have ⟨_, _, H1, H2⟩ := TY.app_inv H' 195 | have ⟨m2l, hl⟩ := ih _ H1 e_ih.1.2 _ _ hm1 196 | have ⟨e', hs⟩ := e_ih.2.1 H2 197 | exact ⟨Option.rec e' m2l, Option.rec hs hl⟩ 198 | · exact hn ⟨ 199 | fun _ _ _ hn => h ⟨_, _, _, hn⟩, 200 | fun _ _ _ _ h1 h2 h3 => h' ⟨_, _, _, _, h1, h2, h3⟩⟩ 201 | cases e with 202 | | bvar i => exact ⟨_, .bvar⟩ 203 | | sort => exact ⟨_, .sort⟩ 204 | | const n ls => exact neut _ H e_ih fun hn => ⟨_, .const hn⟩ 205 | | app ih1 ih2 => 206 | have ⟨_, _, hf, ha⟩ := TY.app_inv H 207 | have ⟨_, h1⟩ := e_ih.1.1 hf 208 | have ⟨_, h2⟩ := e_ih.2.1 ha 209 | exact neut _ H e_ih fun hn => ⟨_, .app hn h1 h2⟩ 210 | | lam ih1 ih2 => 211 | have ⟨_, _, hA, he⟩ := TY.lam_inv H 212 | have ⟨_, h1⟩ := e_ih.1.1 hA 213 | have ⟨_, h2⟩ := e_ih.2.1 he 214 | exact ⟨_, .lam h1 h2⟩ 215 | | forallE ih1 ih2 => 216 | have ⟨_, _, hA, hB⟩ := TY.forallE_inv H 217 | have ⟨_, h1⟩ := e_ih.1.1 hA 218 | have ⟨_, h2⟩ := e_ih.2.1 hB 219 | exact ⟨_, .forallE h1 h2⟩ 220 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/TypeChecker/WHNF.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.TypeChecker.Reduce 2 | 3 | namespace Lean4Lean.TypeChecker.Inner 4 | open Lean hiding Environment Exception 5 | 6 | theorem reduceRecursor.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 7 | RecM.WF c s (reduceRecursor e cheapRec cheapProj) fun oe _ => 8 | ∀ e₁, oe = some e₁ → c.FVarsBelow e e₁ ∧ c.TrExpr e₁ e' := sorry 9 | 10 | theorem whnfFVar.WF {c : VContext} {s : VState} (he : c.TrExprS (.fvar fv) e') : 11 | RecM.WF c s (whnfFVar (.fvar fv) cheapRec cheapProj) fun e₁ _ => 12 | c.FVarsBelow (.fvar fv) e₁ ∧ c.TrExpr e₁ e' := by 13 | refine .getLCtx ?_ 14 | simp [Expr.fvarId!]; split <;> [skip; exact .pure ⟨.rfl, he.trExpr c.Ewf c.Δwf⟩] 15 | rename_i decl h 16 | rw [c.trlctx.1.find?_eq_find?_toList] at h 17 | have := List.find?_some h; simp at this; subst this 18 | let ⟨e', ty', h1, h2, _, h3, _⟩ := 19 | c.trlctx.find?_of_mem c.Ewf (List.mem_of_find?_eq_some h) 20 | refine (whnfCore.WF h3).mono fun _ _ _ ⟨h4, h5⟩ => ?_ 21 | refine ⟨h2.trans h4, h5.defeq c.Ewf c.Δwf ?_⟩ 22 | refine (TrExprS.fvar h1).uniq c.Ewf ?_ he 23 | exact .refl c.Ewf c.Δwf 24 | 25 | theorem reduceProj.WF {c : VContext} {s : VState} (he : c.TrExprS (.proj n i e) e') : 26 | RecM.WF c s (reduceProj i e cheapRec cheapProj) fun oe _ => 27 | ∀ e₁, oe = some e₁ → c.FVarsBelow (.proj n i e) e₁ ∧ c.TrExpr e₁ e' := sorry 28 | 29 | theorem whnfCore'.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 30 | RecM.WF c s (whnfCore' e cheapRec cheapProj) fun e₁ _ => 31 | c.FVarsBelow e e₁ ∧ c.TrExpr e₁ e' := by 32 | unfold whnfCore'; extract_lets F G 33 | let full := (· matches Expr.fvar _ | .app .. | .letE .. | .proj ..) 34 | generalize hP : (fun e₁ (_ : VState) => _) = P 35 | have hid {s} : RecM.WF c s (pure e) P := hP ▸ .pure ⟨.rfl, he.trExpr c.Ewf c.Δwf⟩ 36 | suffices hG : full e → RecM.WF c s (G ⟨⟩) P by 37 | split 38 | any_goals exact hid 39 | any_goals exact hG rfl 40 | · let .mdata he := he 41 | exact (whnfCore'.WF he).bind fun _ _ _ h => hP ▸ .pure h 42 | · refine .getLCtx ?_; split <;> [exact hid; exact hG rfl] 43 | simp [G]; refine fun hfull => .get ?_; split 44 | · rename_i r eq; refine .stateWF fun wf => hP ▸ .pure ?_ 45 | have ⟨_, h1, h2, h3⟩ := (wf.whnfCore_wf eq).2.2.2.2 he.fvarsIn 46 | refine ⟨h1, h3.defeq c.Ewf c.Δwf ?_⟩ 47 | exact h2.uniq c.Ewf (.refl c.Ewf c.Δwf) he 48 | have hsave {e₁ s} (h1 : c.FVarsBelow e e₁) (h2 : c.TrExpr e₁ e') : 49 | (save e cheapRec cheapProj e₁).WF c s P := by 50 | simp [save] 51 | split <;> [skip; exact hP ▸ .pure ⟨h1, h2⟩] 52 | rintro _ mwf wf a s' ⟨⟩ 53 | refine let s' := _; ⟨s', rfl, ?_⟩ 54 | have hic {ic} (hic : WHNFCache.WF c s ic) : WHNFCache.WF c s (ic.insert e e₁) := by 55 | intro _ _ h 56 | rw [Std.HashMap.getElem?_insert] at h; split at h <;> [cases h; exact hic h] 57 | rename_i eq 58 | refine .mk c.mlctx.noBV (.eqv h1 eq BEq.rfl) (he.eqv eq) h2 (.eqv eq ?_) ?_ --_ (.eqv h2 eq BEq.rfl) (.eqv eq ?_) ?_ 59 | · exact he.fvarsIn.mono wf.ngen_wf 60 | · exact h2.fvarsIn.mono wf.ngen_wf 61 | exact hP ▸ ⟨.rfl, { wf with whnfCore_wf := hic wf.whnfCore_wf }, h1, h2⟩ 62 | unfold F; split <;> cases hfull 63 | · simp; exact hP ▸ whnfFVar.WF he 64 | · rename_i fn arg _; generalize eq : fn.app arg = e at * 65 | rw [Expr.withRevApp_eq] 66 | have ⟨_, stk⟩ := AppStack.build <| e.mkAppList_getAppArgsList ▸ he 67 | refine (whnfCore.WF stk.tr).bind fun _ s _ ⟨h1, h2⟩ => ?_ 68 | split <;> [rename_i name dom body bi _; split] 69 | · let rec loop.WF {e e' i rargs f} (H : LambdaBodyN i e' f) (hi : i ≤ rargs.size) : 70 | ∃ n f', LambdaBodyN n e' f' ∧ n ≤ rargs.size ∧ 71 | loop e cheapRec cheapProj rargs i f = loop.cont e cheapRec cheapProj rargs n f' := by 72 | unfold loop; split 73 | · split 74 | · refine loop.WF (by simpa [Nat.add_comm] using H.add (.succ .zero)) ‹_› 75 | · exact ⟨_, _, H, hi, rfl⟩ 76 | · exact ⟨_, _, H, hi, rfl⟩ 77 | refine 78 | let ⟨i, f, h3, h4, eq⟩ := loop.WF (e' := .lam name dom body bi) (.succ .zero) <| by 79 | simp [← eq, Expr.getAppRevArgs_eq, Expr.getAppArgsRevList] 80 | eq ▸ ?_; clear eq 81 | simp [Expr.getAppRevArgs_eq] at h4 ⊢ 82 | obtain ⟨l₁, l₂, h5, rfl⟩ : ∃ l₁ l₂, e.getAppArgsRevList = l₁ ++ l₂ ∧ l₂.length = i := 83 | ⟨_, _, (List.take_append_drop (e.getAppArgsRevList.length - i) ..).symm, by simp; omega⟩ 84 | simp [loop.cont, h5, List.take_of_length_le] 85 | rw [Expr.mkAppRevRange_eq_rev (l₁ := []) (l₂ := l₁) (l₃ := l₂) (by simp) (by rfl) (by rfl)] 86 | have br := BetaReduce.inst_reduce (l₁ := l₂.reverse) 87 | [] (by simpa using h3) (Expr.instantiateList_append ..) (h := by 88 | have := h5 ▸ (c.mlctx.noBV ▸ he.closed).getAppArgsRevList 89 | simp [or_imp, forall_and] at this ⊢ 90 | exact this.2) |>.mkAppRevList (es := l₁) 91 | simp [← Expr.mkAppRevList_reverse, ← Expr.mkAppRevList_append, ← h5] at br 92 | have := h2.rebuild_mkAppRevList c.Ewf c.Δwf stk.tr <| 93 | e.mkAppRevList_getAppArgsRevList ▸ he 94 | have ⟨_, a1, a2⟩ := this.beta c.Ewf c.Δwf br 95 | refine (whnfCore.WF a1).bind fun _ _ _ ⟨b1, b2⟩ => ?_ 96 | have hb := e.mkAppRevList_getAppArgsRevList ▸ h1.mkAppRevList 97 | exact hsave (hb.trans (.betaReduce br) |>.trans b1) <| 98 | b2.defeq c.Ewf c.Δwf a2 99 | · refine (reduceRecursor.WF he).bind fun _ _ _ h => ?_ 100 | split <;> [skip; exact hid] 101 | let ⟨h1, _, h2, eq⟩ := h _ rfl 102 | refine hP ▸ (whnfCore.WF h2).mono fun _ _ _ ⟨h3, h4⟩ => ?_ 103 | exact ⟨h1.trans h3, h4.defeq c.Ewf c.Δwf eq⟩ 104 | · rw [Expr.mkAppRevRange_eq_rev (l₁ := []) (l₃ := []) 105 | (by simp [Expr.getAppRevArgs_toList]; rfl) (by rfl) (by simp [Expr.getAppRevArgs_eq])] 106 | have {e e₁ : Expr} (hb : c.FVarsBelow e e₁) {es e₀' e'} 107 | (hes : c.TrExprS (e.mkAppRevList es) e₀') (he : c.TrExprS e e') (he₁ : c.TrExpr e₁ e') : 108 | c.FVarsBelow (e.mkAppRevList es) (e₁.mkAppRevList es) ∧ 109 | c.TrExpr (e₁.mkAppRevList es) e₀' := by 110 | induction es generalizing e₁ e₀' e' with 111 | | nil => 112 | refine ⟨hb, he₁.defeq c.Ewf c.Δwf ?_⟩ 113 | exact he.uniq c.Ewf (.refl c.Ewf c.Δwf) hes 114 | | cons _ _ ih => 115 | have .app h1 h2 h3 h4 := hes 116 | have ⟨h5, h6⟩ := ih hb h3 he he₁ 117 | exact ⟨fun _ hP he => ⟨h5 _ hP he.1, he.2⟩, 118 | .app c.Ewf c.Δwf h1 h2 h6 (h4.trExpr c.Ewf c.Δwf)⟩ 119 | have eq := e.mkAppRevList_getAppArgsRevList 120 | let ⟨h3, _, h4, eq⟩ := eq ▸ this h1 (eq ▸ he) stk.tr h2 121 | refine (whnfCore.WF h4).bind fun _ _ _ ⟨h5, h6⟩ => ?_ 122 | refine hsave (h3.trans h5) (h6.defeq c.Ewf c.Δwf eq) 123 | · let .letE h1 h2 h3 h4 := he; simp 124 | refine (whnfCore.WF (h4.inst_let c.Ewf.ordered h3)).bind fun _ _ _ ⟨h1, h2⟩ => ?_ 125 | exact hsave (.trans (fun _ _ he => he.2.2.instantiate1 he.2.1) h1) h2 126 | · refine (reduceProj.WF he).bind fun _ _ _ H => ?_ 127 | split 128 | · let ⟨h1, _, h2, eq⟩ := H _ rfl 129 | refine (whnfCore.WF h2).bind fun _ _ _ ⟨h3, h4⟩ => ?_ 130 | exact hsave (h1.trans h3) (h4.defeq c.Ewf c.Δwf eq) 131 | · exact hsave .rfl (he.trExpr c.Ewf c.Δwf) 132 | 133 | theorem whnf'.WF {c : VContext} {s : VState} (he : c.TrExprS e e') : 134 | RecM.WF c s (whnf' e) fun e₁ _ => c.FVarsBelow e e₁ ∧ c.TrExpr e₁ e' := by 135 | unfold whnf'; extract_lets F G 136 | generalize hP : (fun e₁ (_ : VState) => _) = P 137 | have hid {s} : RecM.WF c s (pure e) P := hP ▸ .pure ⟨.rfl, he.trExpr c.Ewf c.Δwf⟩ 138 | suffices hG : RecM.WF c s (G ()) P by 139 | split 140 | any_goals exact hid 141 | any_goals exact hG 142 | · let .mdata he := he 143 | exact (whnf'.WF he).bind fun _ _ _ h => hP ▸ .pure h 144 | · refine .getLCtx ?_; split <;> [exact hid; exact hG] 145 | simp [G]; refine .get ?_; split 146 | · rename_i r eq; refine .stateWF fun wf => hP ▸ .pure ?_ 147 | have ⟨_, h1, h2, h3⟩ := (wf.whnf_wf eq).2.2.2.2 he.fvarsIn 148 | refine ⟨h1, h3.defeq c.Ewf c.Δwf ?_⟩ 149 | exact h2.uniq c.Ewf (.refl c.Ewf c.Δwf) he 150 | unfold F 151 | have {e e' s n} (he : c.TrExprS e e') : (loop e n).WF c s fun e₁ _ => 152 | c.FVarsBelow e e₁ ∧ c.TrExpr e₁ e' := by 153 | induction n generalizing s e e' with | zero => exact .throw | succ n ih => ?_ 154 | refine .getEnv <| (whnfCore'.WF he).bind fun e₁ s _ ⟨h1, _, he₁, eq⟩ => ?_ 155 | refine (M.WF.liftExcept reduceNative.WF).lift.bind fun _ _ _ h3 => ?_ 156 | extract_lets F1 F2; split <;> [cases h3 _ rfl; skip] 157 | refine .pureBind ?_; unfold F2 158 | refine (reduceNat.WF he₁).bind fun _ _ _ h3 => ?_; split 159 | · exact .pure ⟨.trans h1 (h3 _ rfl).1, (h3 _ rfl).2.defeq c.Ewf c.Δwf eq⟩ 160 | refine .pureBind ?_; unfold F1; split <;> [skip; exact .pure ⟨h1, _, he₁, eq⟩] 161 | have ⟨a1, _, a2, eq'⟩ := (unfoldDefinition.WF he₁ ‹_› :) 162 | refine (ih a2).mono fun _ _ _ ⟨b1, b2⟩ => ?_ 163 | exact ⟨h1.trans <| a1.trans b1, b2.defeq c.Ewf c.Δwf <| eq'.trans c.Ewf c.Δwf eq⟩ 164 | refine (this he).bind fun e₁ s _ ⟨h1, h2⟩ => ?_ 165 | rintro _ mwf wf a s' ⟨⟩ 166 | refine let s' := _; ⟨s', rfl, ?_⟩ 167 | have hic {ic} (hic : WHNFCache.WF c s ic) : WHNFCache.WF c s (ic.insert e e₁) := by 168 | intro _ _ h 169 | rw [Std.HashMap.getElem?_insert] at h; split at h <;> [cases h; exact hic h] 170 | rename_i eq 171 | refine .mk c.mlctx.noBV (.eqv h1 eq BEq.rfl) (he.eqv eq) h2 (.eqv eq ?_) ?_ 172 | · exact he.fvarsIn.mono wf.ngen_wf 173 | · exact h2.fvarsIn.mono wf.ngen_wf 174 | exact hP ▸ ⟨.rfl, { wf with whnf_wf := hic wf.whnf_wf }, h1, h2⟩ 175 | -------------------------------------------------------------------------------- /Lean4Lean/Std/Basic.lean: -------------------------------------------------------------------------------- 1 | import Batteries.CodeAction 2 | import Batteries.Data.Array.Lemmas 3 | import Batteries.Data.HashMap.Basic 4 | import Batteries.Data.UnionFind.Basic 5 | import Batteries.Tactic.SeqFocus 6 | 7 | open Std 8 | 9 | attribute [simp] Option.bind_eq_some List.filterMap_cons 10 | 11 | protected theorem Nat.le_iff_exists_add {a b : Nat} : a ≤ b ↔ ∃ c, b = a + c := 12 | ⟨fun h => ⟨_, (Nat.add_sub_cancel' h).symm⟩, fun ⟨_, h⟩ => h ▸ Nat.le_add_right ..⟩ 13 | 14 | protected theorem Nat.le_iff_exists_add' {a b : Nat} : a ≤ b ↔ ∃ c, b = c + a := by 15 | simp [Nat.add_comm, Nat.le_iff_exists_add] 16 | 17 | protected theorem List.Forall₂.rfl 18 | {R : α → α → Prop} {l : List α} (h : ∀ a ∈ l, R a a) : l.Forall₂ R l := by 19 | induction l with 20 | | nil => constructor 21 | | cons _ _ ih => simp at h; exact .cons h.1 (ih h.2) 22 | 23 | @[simp] theorem List.forall₂_nil_left_iff {l} : Forall₂ R nil l ↔ l = nil := 24 | ⟨fun H => by cases H; rfl, by rintro rfl; exact Forall₂.nil⟩ 25 | 26 | @[simp] theorem List.forall₂_nil_right_iff {l} : Forall₂ R l nil ↔ l = nil := 27 | ⟨fun H => by cases H; rfl, by rintro rfl; exact Forall₂.nil⟩ 28 | 29 | theorem List.forall₂_cons_left_iff {a l u} : 30 | Forall₂ R (a :: l) u ↔ ∃ b u', R a b ∧ Forall₂ R l u' ∧ u = b :: u' := 31 | Iff.intro 32 | (fun h => match u, h with | b :: u', Forall₂.cons h₁ h₂ => ⟨b, u', h₁, h₂, rfl⟩) 33 | (fun h => match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂) 34 | 35 | theorem List.forall₂_cons_right_iff {b l u} : 36 | Forall₂ R u (b :: l) ↔ ∃ a u', R a b ∧ Forall₂ R u' l ∧ u = a :: u' := 37 | Iff.intro 38 | (fun h => match u, h with | b :: u', Forall₂.cons h₁ h₂ => ⟨b, u', h₁, h₂, rfl⟩) 39 | (fun h => match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂) 40 | 41 | theorem List.Forall₂.imp (H : ∀ a b, R a b → S a b) 42 | {l₁ l₂} (h : Forall₂ R l₁ l₂) : Forall₂ S l₁ l₂ := by 43 | induction h <;> constructor <;> [(apply H; assumption); assumption] 44 | 45 | theorem List.Forall₂.trans (H : ∀ a b c, R a b → S b c → T a c) 46 | {l₁ l₂ l₃} (h₁ : Forall₂ R l₁ l₂) (h₂ : Forall₂ S l₂ l₃) : Forall₂ T l₁ l₃ := by 47 | induction h₁ generalizing l₃ <;> cases h₂ <;> constructor <;> solve_by_elim 48 | 49 | theorem List.Forall₂.and {l₁ l₂} (h₁ : Forall₂ R l₁ l₂) (h₂ : Forall₂ S l₁ l₂) : 50 | Forall₂ (fun x y => R x y ∧ S x y) l₁ l₂ := by induction h₁ <;> simp_all 51 | 52 | @[simp] theorem List.forall₂_map_left_iff {f : γ → α} : 53 | ∀ {l u}, Forall₂ R (map f l) u ↔ Forall₂ (fun c b => R (f c) b) l u 54 | | [], _ => by simp only [map, forall₂_nil_left_iff] 55 | | a :: l, _ => by simp only [map, forall₂_cons_left_iff, forall₂_map_left_iff] 56 | 57 | @[simp] theorem List.forall₂_map_right_iff {f : γ → β} : 58 | ∀ {l u}, Forall₂ R l (map f u) ↔ Forall₂ (fun a c => R a (f c)) l u 59 | | _, [] => by simp only [map, forall₂_nil_right_iff] 60 | | _, b :: u => by simp only [map, forall₂_cons_right_iff, forall₂_map_right_iff] 61 | 62 | theorem List.Forall₂.flip : ∀ {a b}, Forall₂ (flip R) b a → Forall₂ R a b 63 | | _, _, Forall₂.nil => Forall₂.nil 64 | | _ :: _, _ :: _, Forall₂.cons h₁ h₂ => Forall₂.cons h₁ h₂.flip 65 | 66 | theorem List.Forall₂.forall_exists_l {l₁ l₂} (h : Forall₂ R l₁ l₂) : ∀ a ∈ l₁, ∃ b ∈ l₂, R a b := by 67 | induction h with simp [*] | cons _ _ ih => exact fun a h => .inr (ih _ h) 68 | 69 | theorem List.Forall₂.forall_exists_r {l₁ l₂} (h : Forall₂ R l₁ l₂) : ∀ b ∈ l₂, ∃ a ∈ l₁, R a b := 70 | h.flip.forall_exists_l 71 | 72 | theorem List.Forall₂.length_eq : ∀ {l₁ l₂}, Forall₂ R l₁ l₂ → length l₁ = length l₂ 73 | | _, _, Forall₂.nil => rfl 74 | | _, _, Forall₂.cons _ h₂ => congrArg Nat.succ (Forall₂.length_eq h₂) 75 | 76 | theorem List.Forall₂.append_of_left : ∀ {l₁ l₂ r₁ r₂}, length l₁ = length l₂ → 77 | (Forall₂ R (l₁ ++ r₁) (l₂ ++ r₂) ↔ Forall₂ R l₁ l₂ ∧ Forall₂ R r₁ r₂) 78 | | [], [], _, _, _ => by simp 79 | | a::l₁, b::l₂, _, _, eq => by simp [append_of_left (Nat.succ_inj.1 eq), and_assoc] 80 | 81 | theorem List.Forall₂.append_of_right {l₁ l₂ r₁ r₂} (H : length r₁ = length r₂) : 82 | Forall₂ R (l₁ ++ r₁) (l₂ ++ r₂) ↔ Forall₂ R l₁ l₂ ∧ Forall₂ R r₁ r₂ := by 83 | refine ⟨fun h => (append_of_left ?_).1 h, fun h => (append_of_left h.1.length_eq).2 h⟩ 84 | simpa [H] using h.length_eq 85 | 86 | theorem List.map_id''' {f : α → α} (l : List α) (h : ∀ x ∈ l, f x = x) : map f l = l := by 87 | induction l <;> simp_all 88 | 89 | theorem List.map_fst_lookup {f : α → β} [BEq β] (l : List α) (b : β) : 90 | (l.map (fun a => (f a, a))).lookup b = l.find? fun a => b == f a := by 91 | induction l <;> simp_all [lookup, find?] 92 | 93 | def List.All (P : α → Prop) : List α → Prop 94 | | [] => True 95 | | a::as => P a ∧ as.All P 96 | 97 | theorem List.All.imp {P Q : α → Prop} (h : ∀ a, P a → Q a) : ∀ {l : List α}, l.All P → l.All Q 98 | | [] => id 99 | | _::_ => And.imp (h _) (List.All.imp h) 100 | 101 | theorem List.append_eq_append_of_length_le {a b c d : List α} (h : length a ≤ length c) : 102 | a ++ b = c ++ d ↔ ∃ a', c = a ++ a' ∧ b = a' ++ d := by 103 | rw [append_eq_append_iff, or_iff_left_iff_imp] 104 | rintro ⟨c', rfl, rfl⟩ 105 | rw [← Nat.add_zero c.length, length_append, 106 | Nat.add_le_add_iff_left, Nat.le_zero, length_eq_zero_iff] at h 107 | subst h; exact ⟨[], by simp⟩ 108 | 109 | @[simp] theorem List.nodup_reverse {l : List α} : Nodup (reverse l) ↔ Nodup l := 110 | pairwise_reverse.trans <| by simp only [Nodup, Ne, eq_comm] 111 | 112 | theorem List.foldl_congr 113 | (H : ∀ a, ∀ x ∈ l, f a x = g a x) : foldl f a l = foldl g a l := by 114 | induction l generalizing a <;> simp_all 115 | 116 | theorem List.idxOf_eq_length_iff [BEq α] [LawfulBEq α] 117 | {a : α} {l : List α} : idxOf a l = length l ↔ a ∉ l := by 118 | induction l with 119 | | nil => exact iff_of_true rfl not_mem_nil 120 | | cons b l ih => 121 | simp only [length, mem_cons, idxOf_cons] 122 | rw [cond_eq_if] 123 | split <;> rename_i h <;> simp at h 124 | · exact iff_of_false (by rintro ⟨⟩) fun H => H <| Or.inl h.symm 125 | · simp only [Ne.symm h, false_or] 126 | rw [← ih] 127 | exact Nat.succ_inj 128 | 129 | instance [BEq α] [LawfulBEq α] : PartialEquivBEq α where 130 | symm h := by simp at *; exact h.symm 131 | trans h1 h2 := by simp at *; exact h1.trans h2 132 | 133 | theorem beq_comm [BEq α] [PartialEquivBEq α] (a b : α) : (a == b) = (b == a) := 134 | Bool.eq_iff_iff.2 ⟨PartialEquivBEq.symm, PartialEquivBEq.symm⟩ 135 | 136 | theorem List.mapM_eq_some {f : α → Option β} {l : List α} {l' : List β} : 137 | l.mapM f = some l' ↔ List.Forall₂ (f · = some ·) l l' := by 138 | induction l generalizing l' with 139 | | nil => simp only [mapM_nil, pure, Option.some.injEq, forall₂_nil_left_iff, @eq_comm _ l'] 140 | | cons x l ih => 141 | simp [mapM_cons, Bind.bind, pure, Option.bind_eq_some_iff, Option.some.injEq, 142 | forall₂_cons_left_iff, @eq_comm _ l', exists_and_left, ih] 143 | 144 | @[simp] theorem Option.bind_eq_none'' {o : Option α} {f : α → Option β} : 145 | o.bind f = none ↔ ∀ a, o = some a → f a = none := by cases o <;> simp 146 | 147 | @[simp] theorem Option.forall_ne_some {o : Option α} : (∀ a, o ≠ some a) ↔ o = none := by 148 | cases o <;> simp 149 | 150 | @[simp] theorem Option.orElse_eq_none {a : Option α} {b : Unit → Option α} : 151 | a.orElse b = none ↔ a = none ∧ b () = none := by 152 | cases a <;> simp [Option.orElse] 153 | 154 | instance [BEq α] [PartialEquivBEq α] [BEq β] [PartialEquivBEq β] : PartialEquivBEq (α × β) where 155 | symm := by simp [(· == ·)]; grind [BEq.symm] 156 | trans := by simp [(· == ·)]; grind [BEq.trans] 157 | 158 | instance [BEq α] [EquivBEq α] [BEq β] [EquivBEq β] : EquivBEq (α × β) where 159 | rfl := by simp [(· == ·)] 160 | 161 | instance [BEq α] [PartialEquivBEq α] : PartialEquivBEq (List α) where 162 | symm := by 163 | simp [(· == ·)]; intro a b 164 | induction a generalizing b <;> cases b <;> simp [List.beq]; grind [BEq.symm] 165 | trans := by 166 | simp [(· == ·)]; intro a b c 167 | induction a generalizing b c <;> cases b <;> simp [List.beq] 168 | cases c <;> simp [List.beq]; grind [BEq.trans] 169 | 170 | instance [BEq α] [EquivBEq α] : EquivBEq (List α) where 171 | rfl {a} := by simp [(· == ·)]; induction a <;> simp [List.beq, *] 172 | 173 | namespace BitVec 174 | 175 | variable (n : Nat) 176 | 177 | instance : TransOrd (BitVec n) := 178 | TransOrd.compareOfLessAndEq_of_antisymm_of_trans_of_total_of_not_le 179 | BitVec.le_antisymm BitVec.le_trans BitVec.le_total BitVec.not_le 180 | 181 | instance : LawfulEqOrd (BitVec n) where 182 | eq_of_compare := compareOfLessAndEq_eq_eq BitVec.le_refl BitVec.not_le |>.mp 183 | 184 | end BitVec 185 | 186 | namespace UInt64 187 | 188 | variable (n : Nat) 189 | 190 | instance : TransOrd UInt64 := 191 | TransOrd.compareOfLessAndEq_of_antisymm_of_trans_of_total_of_not_le 192 | UInt64.le_antisymm UInt64.le_trans UInt64.le_total UInt64.not_le 193 | 194 | instance : LawfulEqOrd UInt64 where 195 | eq_of_compare := compareOfLessAndEq_eq_eq UInt64.le_refl UInt64.not_le |>.mp 196 | 197 | end UInt64 198 | 199 | namespace Batteries.UnionFind 200 | 201 | @[simp] theorem size_empty : (∅ : UnionFind).size = 0 := rfl 202 | 203 | @[simp] theorem size_push (uf : UnionFind) : uf.push.size = uf.size + 1 := by 204 | simp [push, size] 205 | 206 | @[simp] theorem size_link (uf : UnionFind) (i j hi) : (uf.link i j hi).size = uf.size := by 207 | simp [link, size] 208 | 209 | @[simp] theorem size_union (uf : UnionFind) (i j) : (uf.union i j).size = uf.size := by 210 | simp [union, size] 211 | 212 | theorem Equiv.eq_of_ge_size (h : Equiv uf a b) (h2 : uf.size ≤ a) : a = b := by 213 | simp [Equiv, rootD, Nat.not_lt.2 h2] at h; split at h 214 | · have := (uf.root ⟨b, ‹_›⟩).2; omega 215 | · exact h 216 | 217 | theorem Equiv.lt_size (h : Equiv uf a b) : a < uf.size ↔ b < uf.size := 218 | suffices ∀ {a b}, Equiv uf a b → b < uf.size → a < uf.size from ⟨this h.symm, this h⟩ 219 | fun h h1 => Nat.not_le.1 fun h2 => Nat.not_le.2 h1 <| h.eq_of_ge_size h2 ▸ h2 220 | 221 | 222 | end Batteries.UnionFind 223 | -------------------------------------------------------------------------------- /Lean4Lean/Level.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Lean4Lean.List 3 | 4 | namespace Lean.Level 5 | 6 | def forEach [Monad m] (l : Level) (f : Level → m Bool) : m Unit := do 7 | if !(← f l) then return 8 | match l with 9 | | .succ l => l.forEach f 10 | | .max l₁ l₂ | .imax l₁ l₂ => l₁.forEach f; l₂.forEach f 11 | | .zero | .param .. | .mvar .. => pure () 12 | 13 | def getUndefParam (l : Level) (ps : List Name) : Option Name := Id.run do 14 | (·.2) <$> StateT.run (s := none) do 15 | l.forEach fun l => do 16 | if !l.hasParam || (← get).isSome then 17 | return false 18 | if let .param n := l then 19 | if n ∉ ps then 20 | set (some n) 21 | return true 22 | 23 | /-! 24 | ## Level normalization 25 | 26 | Based on Yoan Géran, "A Canonical Form for Universe Levels in Impredicative Type Theory" 27 | . 28 | -/ 29 | 30 | namespace Normalize 31 | 32 | local instance : Ord Name := ⟨Name.cmp⟩ 33 | 34 | structure VarNode where 35 | var : Name 36 | offset : Nat 37 | deriving BEq, Ord, Repr 38 | 39 | structure Node where 40 | path : List Name := [] 41 | const : Nat := 0 42 | var : List VarNode := [] 43 | deriving Repr, Inhabited 44 | 45 | instance : BEq Node where 46 | beq n₁ n₂ := n₁.const == n₂.const && n₁.var == n₂.var 47 | instance : Ord Node where 48 | compare n₁ n₂ := compare n₁.const n₂.const |>.then <| compare n₁.var n₂.var 49 | 50 | def subset (cmp : α → α → Ordering) : List α → List α → Bool 51 | | [], _ => true 52 | | _, [] => false 53 | | x :: xs, y :: ys => 54 | match cmp x y with 55 | | .lt => false 56 | | .eq => subset cmp xs ys 57 | | .gt => subset cmp (x :: xs) ys 58 | 59 | def orderedInsert (cmp : α → α → Ordering) (a : α) : List α → Option (List α) 60 | | [] => some [a] 61 | | b :: l => 62 | match cmp a b with 63 | | .lt => some (a :: b :: l) 64 | | .eq => none 65 | | .gt => (orderedInsert cmp a l).map (b :: ·) 66 | 67 | def NormLevel := Std.TreeMap (List Name) Node compare 68 | deriving Repr 69 | 70 | instance : BEq NormLevel where 71 | beq l₁ l₂ := 72 | (l₁.all fun p n => l₂.get? p == some n) && 73 | (l₂.all fun p n => l₁.get? p == some n) 74 | 75 | def VarNode.addVar (v : Name) (k : Nat) : List VarNode → List VarNode 76 | | [] => [⟨v, k⟩] 77 | | v' :: l => 78 | match Name.cmp v v'.var with 79 | | .lt => ⟨v, k⟩ :: v' :: l 80 | | .eq => ⟨v, v'.offset.max k⟩ :: l 81 | | .gt => v' :: addVar v k l 82 | 83 | def NormLevel.addVar (v : Name) (k : Nat) (path' : List Name) (s : NormLevel) : NormLevel := 84 | s.modify path' fun n => { n with var := VarNode.addVar v k n.var } 85 | 86 | def NormLevel.addNode (v : Name) (path' : List Name) (s : NormLevel) : NormLevel := 87 | s.alter path' fun 88 | | none => some { var := [⟨v, 0⟩] } 89 | | some n => some { n with var := VarNode.addVar v 0 n.var } 90 | 91 | def NormLevel.addConst (k : Nat) (path : List Name) (acc : NormLevel) : NormLevel := 92 | if k = 0 || k = 1 && !path.isEmpty then acc else 93 | acc.modify path fun n => { n with const := k.max n.const } 94 | 95 | def normalizeAux (l : Level) (path : List Name) (k : Nat) (acc : NormLevel) : NormLevel := 96 | match l with 97 | | .zero | .imax _ .zero => acc.addConst k path 98 | | .succ u => normalizeAux u path (k+1) acc 99 | | .max u v => normalizeAux u path k acc |> normalizeAux v path k 100 | | .imax u (.succ v) => normalizeAux u path k acc |> normalizeAux v path (k+1) 101 | | .imax u (.max v w) => 102 | normalizeAux (.imax u v) path k acc |> normalizeAux (.imax u w) path k 103 | | .imax u (.imax v w) => 104 | normalizeAux (.imax u w) path k acc |> normalizeAux (.imax v w) path k 105 | | .imax u (.param v) => 106 | match orderedInsert Name.cmp v path with 107 | | some path' => acc.addNode v path' |> normalizeAux u path' k 108 | | none => normalizeAux u path k acc 109 | | .mvar _ | .imax _ (.mvar _) => acc -- unreachable 110 | | .param v => 111 | match orderedInsert Name.cmp v path with 112 | | some path' => 113 | let acc := acc.addConst k path |>.addNode v path' 114 | if k = 0 then acc else acc.addVar v k path' 115 | | none => if k = 0 then acc else acc.addVar v k path 116 | 117 | def subsumeVars : List VarNode → List VarNode → List VarNode 118 | | [], _ => [] 119 | | xs, [] => xs 120 | | x :: xs, y :: ys => 121 | match Name.cmp x.var y.var with 122 | | .lt => x :: subsumeVars xs (y :: ys) 123 | | .eq => if x.offset ≤ y.offset then subsumeVars xs ys else x :: subsumeVars xs ys 124 | | .gt => subsumeVars (x :: xs) ys 125 | 126 | def findParent (f : List Name → Bool) : (l₁ l₂ : List Name) → List Name 127 | | _, [] => [] 128 | | l₁, a :: l₂ => if f (l₁.reverseAux l₂) then [a] else findParent f (a :: l₁) l₂ 129 | 130 | def NormLevel.subsumption (acc : NormLevel) (paths := false) : NormLevel := 131 | acc.foldl (init := acc) fun acc p₁ n₁ => 132 | let n₁ := acc.foldl (init := n₁) fun n₁ p₂ n₂ => 133 | if !subset compare p₂ p₁ then n₁ else 134 | let same := p₁.length == p₂.length 135 | let n₁ := 136 | if n₁.const = 0 || 137 | (same || n₁.const > n₂.const) && 138 | (n₂.var.isEmpty || n₁.const > n₁.var.foldl (·.max ·.offset) 0 + 1) 139 | then n₁ else { n₁ with const := 0 } 140 | if same || n₂.var.isEmpty then n₁ else { n₁ with var := subsumeVars n₁.var n₂.var } 141 | let n₁ := if paths then 142 | let path := findParent acc.contains [] p₁ 143 | let var := if let [v] := path then subsumeVars n₁.var [⟨v, 0⟩] else n₁.var 144 | { n₁ with path, var } 145 | else n₁ 146 | acc.insert p₁ n₁ 147 | 148 | def normalize (l : Level) (paths := false) : NormLevel := 149 | Normalize.normalizeAux l [] 0 (.insert {} [] default) |>.subsumption paths 150 | 151 | def leVars : List VarNode → List VarNode → Bool 152 | | [], _ => true 153 | | _, [] => false 154 | | x :: xs, y :: ys => 155 | match Name.cmp x.var y.var with 156 | | .lt => false 157 | | .eq => x.offset ≤ y.offset && leVars xs ys 158 | | .gt => leVars (x :: xs) ys 159 | 160 | def NormLevel.le (l₁ l₂ : NormLevel) : Bool := 161 | l₁.all fun p₁ n₁ => 162 | if n₁.const = 0 && n₁.var.isEmpty then true else 163 | l₂.any fun p₂ n₂ => 164 | (!n₂.var.isEmpty || n₁.var.isEmpty) && 165 | subset compare p₂ p₁ && 166 | (n₂.const ≤ n₁.const || n₂.var.any (n₂.const ≤ ·.offset + 1)) && 167 | leVars n₁.var n₂.var 168 | 169 | def NormLevel.buildPaths : StateM NormLevel Unit := do 170 | (← get).foldlM (init := ()) fun _ p _ => do 171 | let n := (← get).get! p 172 | if let [v] := n.path then 173 | let l ← getPath (p.erase v) p.length 174 | setPath p (v :: l) 175 | where 176 | setPath (p path : List Name) : StateM NormLevel Unit := 177 | modify (·.modify p ({ · with path })) 178 | 179 | getPath (p : List Name) (depth : Nat) : StateM NormLevel (List Name) := do 180 | let n := (← get).get! p 181 | if let [v] := n.path then 182 | if let depth + 1 := depth then 183 | let l ← getPath (p.erase v) depth 184 | setPath p (v :: l) 185 | return v :: l 186 | return n.path 187 | 188 | structure Tree where 189 | const : Nat 190 | var : List VarNode 191 | child : List (Name × Tree) 192 | deriving Inhabited 193 | 194 | def modifyAt [Inhabited α] (f : α → α) (n : Name) : List (Name × α) → List (Name × α) 195 | | [] => [(n, f default)] 196 | | (x, v) :: l => 197 | match Name.cmp n x with 198 | | .lt => (n, f default) :: (x, v) :: l 199 | | .eq => (x, f v) :: l 200 | | .gt => (x, v) :: modifyAt f n l 201 | 202 | def Tree.modify (path : List Name) (f : Tree → Tree) (t : Tree) : Tree := 203 | match path with 204 | | [] => f t 205 | | a :: p => modify p (t := t) fun t => { t with child := modifyAt f a t.child } 206 | 207 | def NormLevel.toTree (acc : NormLevel) : Tree := 208 | (buildPaths.run acc).run.2.foldl (init := ⟨0, [], []⟩) fun t _ n => 209 | t.modify n.path fun t => { t with const := n.const, var := n.var } 210 | 211 | def treeVarDedup : List VarNode → List (Name × Tree) → List VarNode 212 | | [], _ => [] 213 | | xs, [] => xs 214 | | x :: xs, y :: ys => 215 | match Name.cmp x.1 y.1 with 216 | | .lt => x :: treeVarDedup xs (y :: ys) 217 | | .eq => if x.2 = 0 then treeVarDedup xs ys else x :: treeVarDedup xs ys 218 | | .gt => treeVarDedup (x :: xs) ys 219 | 220 | def Tree.reify : Tree → Level 221 | | { const, var, child } => 222 | let l := child.foldr mkChild none 223 | let l := (treeVarDedup var child).foldr (init := l) fun n r => 224 | some (mkMax (addOffset (.param n.var) n.offset) r) 225 | match l with 226 | | none => ofNat const 227 | | some l => if const = 0 then l else max (ofNat const) l 228 | where 229 | mkMax (l : Level) : Option Level → Level 230 | | none => l 231 | | some u => max l u 232 | mkChild 233 | | (n, t), r => 234 | match reify t with 235 | | .zero => mkMax (.param n) r 236 | | t => mkMax (imax t (.param n)) r 237 | 238 | end Normalize 239 | 240 | def normalize' (l : Level) : Level := (Normalize.normalize l (paths := true)).toTree.reify 241 | 242 | def isEquiv' (u v : Level) : Bool := u == v || Normalize.normalize u == Normalize.normalize v 243 | 244 | def isEquivList : List Level → List Level → Bool := List.all2 isEquiv' 245 | 246 | def geq' (u v : Level) : Bool := (Normalize.normalize v).le (Normalize.normalize u) 247 | 248 | -- local elab "normalize " l:level : command => do 249 | -- Elab.Command.runTermElabM fun _ => do 250 | -- logInfo m!"{normalize' (← Elab.Term.elabLevel l)}" 251 | -- -- logInfo m!"{repr <| Normalize.normalize (← Elab.Term.elabLevel l) }" 252 | 253 | -- universe u v w 254 | -- /-- info: max 1 u -/ 255 | -- #guard_msgs in normalize max u 1 256 | -- /-- info: u -/ 257 | -- #guard_msgs in normalize imax 1 u 258 | -- /-- info: max 1 (imax (u+1) u) -/ 259 | -- #guard_msgs in normalize u+1 260 | -- /-- info: imax 2 u -/ 261 | -- #guard_msgs in normalize imax 2 u 262 | -- /-- info: max v (imax (imax u v) w) -/ 263 | -- #guard_msgs in normalize max w (imax (imax u w) v) 264 | -- /-- info: max v (imax (imax u v) w) -/ 265 | -- #guard_msgs in normalize max (imax (imax u v) w) (imax (imax u w) v) 266 | -- /-- info: u -/ 267 | -- #guard_msgs in normalize imax u u 268 | -- /-- info: max 1 (imax (u+1) u) -/ 269 | -- #guard_msgs in normalize imax u (u+1) 270 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /Lean4Lean/Experimental/StratifiedUntyped.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Lemmas 2 | import Lean4Lean.Theory.Typing.Strong 3 | 4 | namespace Lean4Lean 5 | namespace VEnv 6 | 7 | open VExpr 8 | 9 | section 10 | set_option hygiene false 11 | local notation:65 Γ " ⊢ " e " : " A:30 => HasTypeU1 Γ e A 12 | local notation:65 Γ " ⊢ " e1 " ≡ " e2:30 => IsDefEqU1 Γ e1 e2 13 | 14 | variable (env : VEnv) (uvars : Nat) 15 | 16 | variable (IsDefEqU1 : List VExpr → VExpr → VExpr → Prop) in 17 | inductive HasTypeU1 : List VExpr → VExpr → VExpr → Prop where 18 | | bvar : Lookup Γ i A → Γ ⊢ .bvar i : A 19 | | const : 20 | env.constants c = some (some ci) → 21 | (∀ l ∈ ls, l.WF uvars) → 22 | ls.length = ci.uvars → 23 | Γ ⊢ .const c ls : ci.type.instL ls 24 | | sort : l.WF uvars → Γ ⊢ .sort l : .sort (.succ l) 25 | | app : Γ ⊢ f : .forallE A B → Γ ⊢ a : A → Γ ⊢ .app f a : B.inst a 26 | | lam : Γ ⊢ A : .sort u → A::Γ ⊢ body : B → Γ ⊢ .lam A body : .forallE A B 27 | | forallE : Γ ⊢ A : .sort u → A::Γ ⊢ body : .sort v → Γ ⊢ .forallE A body : .sort (.imax u v) 28 | | defeq : Γ ⊢ A ≡ B → Γ ⊢ e : A → Γ ⊢ e : B 29 | 30 | variable 31 | (HasTypeU1 : List VExpr → VExpr → VExpr → Prop) 32 | (defEq : List VExpr → VExpr → VExpr → Prop) in 33 | inductive IsDefEqU1 : List VExpr → VExpr → VExpr → Prop where 34 | | refl : Γ ⊢ e ≡ e 35 | | symm : Γ ⊢ e ≡ e' → Γ ⊢ e' ≡ e 36 | | trans : Γ ⊢ e₁ ≡ e₂ → Γ ⊢ e₂ ≡ e₃ → Γ ⊢ e₁ ≡ e₃ 37 | | constDF : List.Forall₂ (· ≈ ·) ls ls' → Γ ⊢ .const c ls ≡ .const c ls' 38 | | sortDF : l ≈ l' → Γ ⊢ .sort l ≡ .sort l' 39 | | appDF : 40 | Γ ⊢ f ≡ f' → Γ ⊢ a ≡ a' → Γ ⊢ .app f a ≡ .app f' a' 41 | | lamDF : Γ ⊢ A ≡ A' → A::Γ ⊢ b ≡ b' → Γ ⊢ .lam A b ≡ .lam A' b' 42 | | forallEDF : 43 | Γ ⊢ A ≡ A' → A::Γ ⊢ B ≡ B' → 44 | Γ ⊢ .forallE A B ≡ .forallE A' B' 45 | | beta : A::Γ ⊢ e : B → Γ ⊢ e' : A → Γ ⊢ .app (.lam A e) e' ≡ e.inst e' 46 | | eta : Γ ⊢ .lam A (.app e.lift (.bvar 0)) ≡ e 47 | | proofIrrel : Γ ⊢ p : .sort .zero → Γ ⊢ h : p → Γ ⊢ h' : p → Γ ⊢ h ≡ h' 48 | | extra : 49 | env.defeqs df → (∀ l ∈ ls, l.WF uvars) → ls.length = df.uvars → 50 | Γ ⊢ df.lhs.instL ls ≡ df.rhs.instL ls 51 | 52 | end 53 | 54 | variable! (henv : Ordered env) (hΓ : OnCtx Γ (env.IsType U)) in 55 | theorem IsDefEq.inductionU1 56 | (defEq : List VExpr → VExpr → VExpr → Prop) 57 | (hasType : List VExpr → VExpr → VExpr → Prop) 58 | (hty : ∀ {Γ e A}, HasTypeU1 env U defEq Γ e A → hasType Γ e A) 59 | (hdf : ∀ {Γ e1 e2 A1 A2}, 60 | HasTypeU1 env U defEq Γ e1 A1 → HasTypeU1 env U defEq Γ e2 A2 → 61 | IsDefEqU1 env U hasType Γ e1 e2 → defEq Γ e1 e2) 62 | (H : env.IsDefEq U Γ e1 e2 A) : 63 | HasTypeU1 env U defEq Γ e1 A ∧ 64 | HasTypeU1 env U defEq Γ e2 A ∧ 65 | IsDefEqU1 env U hasType Γ e1 e2 := by 66 | have H' := H.strong henv hΓ; clear hΓ H 67 | induction H' with 68 | | bvar h => exact ⟨.bvar h, .bvar h, .refl⟩ 69 | | symm _ ih => exact ⟨ih.2.1, ih.1, .symm ih.2.2⟩ 70 | | trans _ _ ih1 ih2 => exact ⟨ih1.1, ih2.2.1, .trans ih1.2.2 ih2.2.2⟩ 71 | | @constDF _ _ ls₁ ls₂ _ _ h1 h2 h3 h4 h5 => 72 | exact ⟨.const h1 h2 h4, .defeq sorry <| .const h1 h3 (h5.length_eq.symm.trans h4), .constDF h5⟩ 73 | | @sortDF l l' _ h1 h2 h3 => 74 | refine ⟨.sort h1, ?_, .sortDF h3⟩ 75 | exact .defeq (hdf (.sort (l := l'.succ) h2) (.sort (l := l.succ) h1) 76 | (.sortDF <| VLevel.succ_congr h3.symm)) (.sort h2) 77 | | appDF _ _ _ _ _ _ _ _ _ ihf iha ihBa => 78 | let ⟨hf, hf', ff⟩ := ihf; let ⟨ha, ha', aa⟩ := iha 79 | exact ⟨.app hf ha, .defeq (hdf ihBa.2.1 ihBa.1 (.symm ihBa.2.2)) (.app hf' ha'), .appDF ff aa⟩ 80 | | lamDF _ _ _ _ _ _ _ ihA ihB ihB' ihb ihb' => 81 | refine ⟨.lam ihA.1 ihb.1, .defeq ?_ <| .lam ihA.2.1 ihb'.2.1, .lamDF ihA.2.2 ihb.2.2⟩ 82 | exact hdf (.forallE ihA.2.1 ihB'.1) (.forallE ihA.1 ihB.1) <| 83 | .symm <| .forallEDF ihA.2.2 ihB.2.2 84 | | forallEDF _ _ _ _ _ ih1 ih2 ih3 => 85 | exact ⟨.forallE ih1.1 ih2.1, .forallE ih1.2.1 ih3.2.1, .forallEDF ih1.2.2 ih2.2.2⟩ 86 | | defeqDF _ _ _ ih1 ih2 => 87 | have h := hdf ih1.1 ih1.2.1 ih1.2.2; exact ⟨.defeq h ih2.1, .defeq h ih2.2.1, ih2.2.2⟩ 88 | | beta _ _ _ _ _ _ _ _ ihA _ ihe ihe' _ ihee => 89 | exact ⟨.app (.lam ihA.1 ihe.1) ihe'.1, ihee.1, .beta (hty ihe.1) (hty ihe'.1)⟩ 90 | | eta _ _ _ _ _ _ _ ihA _ _ ihe ihe' => 91 | have := HasTypeU1.app ihe'.1 (.bvar .zero) 92 | rw [instN_bvar0] at this 93 | exact ⟨.lam ihA.1 this, ihe.1, .eta⟩ 94 | | proofIrrel _ _ _ ih1 ih2 ih3 => 95 | exact ⟨ih2.1, ih3.1, .proofIrrel (hty ih1.1) (hty ih2.1) (hty ih3.1)⟩ 96 | | extra h1 h2 h3 _ _ _ _ _ _ _ _ _ ihl' ihr' => 97 | exact ⟨ihl'.1, ihr'.1, .extra h1 h2 h3⟩ 98 | 99 | variable! (henv : Ordered env) (hΓ : OnCtx Γ (env.IsType U)) 100 | {defEq : List VExpr → VExpr → VExpr → Prop} 101 | (IH : ∀ {Γ e1 e2 A}, env.HasType U Γ e1 A → defEq Γ e1 e2 → env.IsDefEq U Γ e1 e2 A) in 102 | theorem HasTypeU1.induction (H : env.HasTypeU1 U defEq Γ e A) : env.HasType U Γ e A := by 103 | induction H with 104 | | bvar h => exact .bvar h 105 | | const h1 h2 h3 => exact .const h1 h2 h3 106 | | sort h => exact .sort h 107 | | app _ _ ih1 ih2 => exact .app (ih1 hΓ) (ih2 hΓ) 108 | | lam _ _ ih1 ih2 => exact .lam (ih1 hΓ) (ih2 ⟨hΓ, _, ih1 hΓ⟩) 109 | | forallE _ _ ih1 ih2 => exact .forallE (ih1 hΓ) (ih2 ⟨hΓ, _, ih1 hΓ⟩) 110 | | defeq h1 _ ih => 111 | let ⟨_, h⟩ := (ih hΓ).isType henv hΓ 112 | exact (IH h h1).defeq (ih hΓ) 113 | 114 | /- 115 | variable (henv : Ordered env) 116 | {hasType : List VExpr → VExpr → VExpr → Prop} 117 | {defEq : List VExpr → VExpr → VExpr → VExpr → Prop} 118 | (hty : ∀ {Γ e A}, hasType Γ e A → env.HasType U Γ e A) 119 | (hdf : ∀ {Γ e1 e2 A}, defEq Γ e1 e2 A → env.IsDefEq U Γ e1 e2 A) in 120 | theorem IsDefEqU1.induction 121 | (H1 : env.HasType U Γ e1 A) (H2 : env.HasType U Γ e2 A) 122 | (H : env.IsDefEqU1 U hasType Γ e1 e2) : env.IsDefEq U Γ e1 e2 A := by 123 | induction H with 124 | | refl => exact H1 125 | | symm _ ih => exact (ih H2 H1).symm 126 | | trans _ _ ih1 ih2 => exact ih1.trans ih2 127 | | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 128 | | appDF _ _ ih1 ih2 => exact .appDF ih1 ih2 129 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 130 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 131 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 132 | | eta h => exact .eta (hty h) 133 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 134 | | extra h1 h2 h3 => exact .extra h1 h2 h3 135 | -/ 136 | 137 | /- 138 | variable 139 | {U : Nat} 140 | (hasType : List VExpr → VExpr → VExpr → Prop) 141 | (hasType' : List VExpr → VExpr → VExpr → Prop) 142 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) 143 | (refl : ∀ {Γ e A}, hasType Γ e A → hasType' Γ e A) 144 | (sort : ∀ {Γ l l'}, l.WF U → l'.WF U → l ≈ l' → 145 | hasType' Γ (.sort l') (.sort l.succ)) 146 | (app : ∀ {Γ f a A B}, 147 | hasType' Γ f (.forallE A B) → hasType' Γ a A → hasType' Γ (.app f a) (B.inst a)) 148 | in 149 | protected theorem IsDefEqU1.hasType_ind 150 | (H : IsDefEqU1 env U hasType defEq Γ e1 e2 A) : 151 | hasType' Γ e1 A ∧ hasType' Γ e2 A := by 152 | induction H with 153 | | refl h => exact ⟨refl h, refl h⟩ 154 | | symm _ ih => exact ih.symm 155 | | trans _ _ ih1 ih2 => exact ⟨ih1.1, ih2.2⟩ 156 | | sortDF h1 h2 h3 => exact ⟨sort h1 h1 rfl, sort h1 h2 h3⟩ 157 | | appDF _ _ ih1 ih2 => 158 | exact ⟨app ih1.1 ih2.1, app _ _⟩ 159 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 160 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 161 | | defeqDF h1 _ ih => exact .defeqDF (hdf h1) ih 162 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 163 | | eta h => exact .eta (hty h) 164 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 165 | | extra h1 h2 h3 => exact .extra h1 h2 h3 166 | 167 | variable 168 | (hasType : List VExpr → VExpr → VExpr → Prop) 169 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) in 170 | inductive IsDefEqU1 : List VExpr → VExpr → VExpr → VLevel → Prop 171 | | refl : hasType Γ A (.sort u) → IsDefEqU1 Γ A A u 172 | 173 | variable (henv : Ordered env) 174 | {hasType : List VExpr → VExpr → VExpr → Prop} 175 | {defEq : List VExpr → VExpr → VExpr → VExpr → Prop} 176 | (refl : ∀ {Γ e A}, hasType Γ e A → defEq' Γ e e A) 177 | (hdf : ∀ {Γ e1 e2 A}, defEq Γ e1 e2 A → env.IsDefEq U Γ e1 e2 A) in 178 | theorem IsDefEqU1.unique_typing1 179 | (H : env.IsDefEqU1 U hasType defEq Γ e1 e2 A) : 180 | hasType Γ e1 A ∧ hasType Γ e2 A := by 181 | induction H with 182 | | refl h => exact hty h 183 | | symm _ ih => exact ih.symm 184 | | trans _ _ ih1 ih2 => exact ih1.trans ih2 185 | | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 186 | | appDF _ _ ih1 ih2 => exact .appDF ih1 ih2 187 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 188 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 189 | | defeqDF h1 _ ih => exact .defeqDF (hdf h1) ih 190 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 191 | | eta h => exact .eta (hty h) 192 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 193 | | extra h1 h2 h3 => exact .extra h1 h2 h3 194 | 195 | variable (henv : Ordered env) in 196 | theorem HasType1.unique_typing' 197 | (H1 : env.HasType1 U (IsDefEqU1 env U hasType defEq) Γ e A1) 198 | (H2 : env.HasType1 U (IsDefEqU1 env U hasType defEq) Γ e A2) : 199 | ∃ u, env.IsDefEqU1 U hasType defEq Γ A1 A2 (.sort u) := by 200 | generalize eq1 : e = e' at H2 201 | induction H1 generalizing e' A2 with subst eq1 202 | | defeq h1 _ ih => 203 | let ⟨_, ih⟩ := ih _ rfl H2 204 | exact ⟨_, h1.trans _⟩ 205 | done 206 | | _ => ?_ 207 | <;> cases H2 208 | 209 | case bvar.bvar _ => 210 | done 211 | case const.const _ _ _ => 212 | done 213 | case sort.sort _ => 214 | done 215 | case app.app _ _ _ _ => 216 | done 217 | case lam.lam _ _ _ _ => 218 | done 219 | case forallE.forallE _ _ _ _ => 220 | done 221 | case defeq.defeq _ _ _ => 222 | done 223 | _ 224 | -- | bvar h => 225 | -- refine .bvar h.instL 226 | -- | @const _ _ ls' _ h1 h2 h3 => 227 | -- simp [instL, instL_instL] 228 | -- exact .const h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 229 | -- | sort _ _ h3 => 230 | -- exact .sortDF (VLevel.WF.inst hls) (VLevel.WF.inst hls) (VLevel.inst_congr_l h3) 231 | -- | app _ _ ih1 ih2 => exact instL_instN ▸ .appDF ih1 ih2 232 | -- | lam _ _ ih1 ih2 => exact .lamDF ih1 ih2 233 | -- | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 234 | -- | defeqDF _ _ ih1 ih2 => exact .defeqDF ih1 ih2 235 | -- | beta _ _ ih1 ih2 => simpa using .beta ih1 ih2 236 | -- | eta _ ih => simpa [instL] using .eta ih 237 | -- | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel ih1 ih2 ih3 238 | -- | extra h1 h2 h3 => 239 | -- simp [instL, instL_instL] 240 | -- exact .extra h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 241 | -/ 242 | 243 | /- 244 | variable (henv : Ordered env) in 245 | theorem IsDefEq.unique_typing' 246 | (H1 : env.IsDefEq U Γ e1 e2 A1) (H2 : env.IsDefEq U Γ e1 e2 A2) : 247 | ∃ u, env.IsDefEq U Γ A1 A2 (.sort u) := by 248 | generalize eq1 : e1 = e1', eq2 : e2 = e2' at H2 249 | induction H1 generalizing A2 with 250 | | bvar h => 251 | 252 | refine .bvar h.instL 253 | | @const _ _ ls' _ h1 h2 h3 => 254 | simp [instL, instL_instL] 255 | exact .const h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 256 | | symm _ ih => exact .symm ih 257 | | trans _ _ ih1 ih2 => exact .trans ih1 ih2 258 | | sortDF _ _ h3 => 259 | exact .sortDF (VLevel.WF.inst hls) (VLevel.WF.inst hls) (VLevel.inst_congr_l h3) 260 | | appDF _ _ ih1 ih2 => exact instL_instN ▸ .appDF ih1 ih2 261 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 262 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 263 | | defeqDF _ _ ih1 ih2 => exact .defeqDF ih1 ih2 264 | | beta _ _ ih1 ih2 => simpa using .beta ih1 ih2 265 | | eta _ ih => simpa [instL] using .eta ih 266 | | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel ih1 ih2 ih3 267 | | extra h1 h2 h3 => 268 | simp [instL, instL_instL] 269 | exact .extra h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 270 | -/ 271 | 272 | 273 | /- depends on church-rosser 274 | variable {env : VEnv} (henv : env.Ordered) in 275 | theorem IsDefEq.weakN_inv (W : Ctx.LiftN n k Γ Γ') 276 | (H : env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k)) : 277 | env.IsDefEq U Γ e1 e2 A := by 278 | generalize eq1 : e1.liftN n k = e1', eq2 : e2.liftN n k = e2', eqA : A.liftN n k = A' at H 279 | induction H generalizing k e1 e2 A with 280 | | bvar h => 281 | cases eqA; cases e1 <;> cases eq1; cases e2 <;> injection eq2 282 | cases liftVar_inj.1 ‹_›; exact .bvar (h.weakN_inv W) 283 | | @const c ci ls Γ h1 h2 h3 => 284 | cases e1 <;> cases eq1; cases e2 <;> cases eq2 285 | rw [ClosedN.liftN_eq_rev (eqA ▸ (henv.closedC h1).instL) (Nat.zero_le _)] at eqA 286 | exact eqA ▸ .const h1 h2 h3 287 | | symm _ ih => exact .symm (ih W eq2 eq1 eqA) 288 | | trans _ _ ih1 ih2 => sorry 289 | -- | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 290 | -- | appDF _ _ ih1 ih2 => exact liftN_inst_hi .. ▸ .appDF (ih1 W) (ih2 W) 291 | -- | lamDF _ _ ih1 ih2 => exact .lamDF (ih1 W) (ih2 W.succ) 292 | -- | forallEDF _ _ ih1 ih2 => exact .forallEDF (ih1 W) (ih2 W.succ) 293 | -- | defeqDF _ _ ih1 ih2 => exact .defeqDF (ih1 W) (ih2 W) 294 | -- | beta _ _ ih1 ih2 => 295 | -- exact liftN_inst_hi .. ▸ liftN_instN_hi .. ▸ .beta (ih1 W.succ) (ih2 W) 296 | -- | eta _ ih => 297 | -- have := IsDefEq.eta (ih W) 298 | -- simp [liftN]; rwa [← lift_liftN'] 299 | -- | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel (ih1 W) (ih2 W) (ih3 W) 300 | -- | extra h1 h2 h3 => 301 | -- have ⟨⟨hA1, _⟩, hA2, hA3⟩ := henv.closed.2 h1 302 | -- rw [ 303 | -- hA1.instL.liftN_eq (Nat.zero_le _), 304 | -- hA2.instL.liftN_eq (Nat.zero_le _), 305 | -- hA3.instL.liftN_eq (Nat.zero_le _)] 306 | -- exact .extra h1 h2 h3 307 | | _ => sorry 308 | 309 | variable {env : VEnv} (henv : env.Ordered) in 310 | theorem HasType.weakN_inv (W : Ctx.LiftN n k Γ Γ') 311 | (H : env.HasType U Γ' (e.liftN n k) (A.liftN n k)) : 312 | env.HasType U Γ e A := IsDefEq.weakN_inv henv W H 313 | 314 | variable {env : VEnv} (henv : env.Ordered) in 315 | theorem IsType.weakN_inv (W : Ctx.LiftN n k Γ Γ') (H : env.IsType U Γ' (A.liftN n k)) : 316 | env.IsType U Γ A := let ⟨_, h⟩ := H; ⟨_, h.weakN_inv henv W⟩ 317 | -/ 318 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/LocalContext.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Std.PersistentHashMap 2 | import Lean4Lean.Verify.Expr 3 | import Lean4Lean.Verify.Typing.Expr 4 | import Lean4Lean.Verify.Typing.Lemmas 5 | 6 | namespace Lean.LocalContext 7 | 8 | noncomputable def toList (lctx : LocalContext) : List LocalDecl := 9 | lctx.decls.toList'.reverse.filterMap id 10 | 11 | noncomputable def fvars (lctx : LocalContext) : List FVarId := 12 | lctx.toList.map (·.fvarId) 13 | 14 | def mkBindingList1 (isLambda : Bool) (lctx : LocalContext) 15 | (xs : List FVarId) (x : FVarId) (b : Expr) : Expr := 16 | match lctx.find? x with 17 | | some (.cdecl _ _ n ty bi _) => 18 | let ty := ty.abstractList xs 19 | if isLambda then 20 | .lam n ty b bi 21 | else 22 | .forallE n ty b bi 23 | | some (.ldecl _ _ n ty val nonDep _) => 24 | if b.hasLooseBVar' 0 then 25 | let ty := ty.abstractList xs 26 | let val := val.abstractList xs 27 | .letE n ty val b nonDep 28 | else 29 | b.lowerLooseBVars' 1 1 30 | | none => panic! "unknown free variable" 31 | 32 | def mkBindingList (isLambda : Bool) (lctx : LocalContext) (xs : List FVarId) (b : Expr) : Expr := 33 | core (b.abstractList xs) 34 | where 35 | core := go xs.reverse 36 | go : List FVarId → Expr → Expr 37 | | [], b => b 38 | | x :: xs, b => go xs (mkBindingList1 isLambda lctx xs.reverse x b) 39 | 40 | theorem mkBinding_eq : 41 | mkBinding isLambda lctx ⟨xs.map .fvar⟩ b = mkBindingList isLambda lctx xs b := by 42 | simp only [mkBinding, List.getElem_toArray, Expr.abstractRange_eq, Expr.hasLooseBVar_eq, 43 | Expr.abstract_eq, ← Array.take_eq_extract, List.take_toArray, Bool.and_false, 44 | ← List.map_take, List.getElem_map, Expr.lowerLooseBVars_eq] 45 | dsimp only [Array.size] 46 | simp only [List.getElem_eq_getElem?_get, Option.get_eq_getD (fallback := default)] 47 | change Nat.foldRev _ (fun i x => 48 | mkBindingList1 isLambda lctx (xs.take i) (xs[i]?.getD default)) .. = mkBindingList.go .. 49 | rw [List.length_map]; generalize eq : xs.length = n 50 | generalize b.abstractList xs = b 51 | induction n generalizing xs b with 52 | | zero => let [] := xs; simp [mkBindingList.go] 53 | | succ n ih => 54 | obtain rfl | ⟨xs, a, rfl⟩ := List.eq_nil_or_concat xs; · cases eq 55 | simp at eq ⊢; subst eq 56 | simp +contextual only [Nat.le_of_lt, List.take_append_of_le_length, 57 | List.getElem?_append_left, mkBindingList.go, ih]; simp 58 | 59 | theorem mkBindingList1_abstract {xs : List FVarId} 60 | (hx : lctx.find? x = some decl) (nd : (a :: xs).Nodup) : 61 | (mkBindingList1 isLambda lctx xs x b).abstract1 a xs.length = 62 | mkBindingList1 isLambda lctx (a :: xs) x (b.abstract1 a (xs.length + 1)) := by 63 | have (e:_) := Nat.zero_add _ ▸ Expr.abstract1_abstractList' (k := 0) (e := e) nd 64 | simp [mkBindingList1, hx]; cases decl with simp 65 | | cdecl _ _ _ ty => split <;> simp [Expr.abstract1, Expr.abstract1, this] 66 | | ldecl => 67 | have := Expr.abstract1_hasLooseBVar a b (xs.length + 1) 0 68 | simp at this; simp [this]; clear this 69 | split 70 | · simp [Expr.abstract1, Expr.abstract1, this] 71 | · rename_i h; simp at h 72 | rw [Expr.abstract1_lower h (Nat.zero_le _)] 73 | 74 | theorem mkBindingList_core_cons {xs : List FVarId} {b : Expr} 75 | (hx : ∀ x ∈ xs, ∃ decl, lctx.find? x = some decl) (nd : (a :: xs).Nodup) : 76 | mkBindingList.core isLambda lctx (a :: xs) (b.abstract1 a xs.length) = 77 | mkBindingList1 isLambda lctx [] a 78 | ((mkBindingList.core isLambda lctx xs b).abstract1 a) := by 79 | obtain ⟨xs, rfl⟩ : ∃ xs', List.reverse xs' = xs := ⟨_, List.reverse_reverse _⟩ 80 | simp [mkBindingList.core] at * 81 | induction xs generalizing b with 82 | | nil => simp [mkBindingList.go] 83 | | cons c xs ih => 84 | simp at hx nd ih 85 | let ⟨decl, eq⟩ := hx.1 86 | simp [mkBindingList.go] 87 | rw [← xs.length_reverse, ← mkBindingList1_abstract eq (by simp [*])] 88 | simp [ih hx.2 nd.1.2 nd.2.2] 89 | 90 | @[simp] theorem mkBindingList_nil : mkBindingList isLambda lctx [] b = b := rfl 91 | 92 | theorem mkBindingList_cons 93 | (hx : ∀ x ∈ xs, ∃ decl, lctx.find? x = some decl) (nd : (a :: xs).Nodup) : 94 | mkBindingList isLambda lctx (a :: xs) b = 95 | mkBindingList1 isLambda lctx [] a ((mkBindingList isLambda lctx xs b).abstract1 a) := by 96 | simp [mkBindingList] 97 | rw [← Expr.abstract1_abstractList' nd] 98 | rw [Nat.zero_add, mkBindingList_core_cons hx nd] 99 | 100 | theorem mkBindingList_eq_fold 101 | (hx : ∀ x ∈ xs, ∃ decl, lctx.find? x = some decl) (nd : xs.Nodup) : 102 | mkBindingList isLambda lctx xs b = 103 | xs.foldr (fun a e => mkBindingList1 isLambda lctx [] a (e.abstract1 a)) b := by 104 | induction xs <;> simp_all [mkBindingList_cons] 105 | 106 | theorem mkBindingList1_congr (H : lctx₁.find? x = lctx₂.find? x) : 107 | mkBindingList1 isLambda lctx₁ xs x b = mkBindingList1 isLambda lctx₂ xs x b := by 108 | simp [mkBindingList1, H] 109 | 110 | theorem mkBindingList_congr 111 | (H : ∀ x ∈ xs, lctx₁.find? x = lctx₂.find? x) : 112 | mkBindingList isLambda lctx₁ xs b = mkBindingList isLambda lctx₂ xs b := by 113 | obtain ⟨xs, rfl⟩ : ∃ xs', List.reverse xs' = xs := ⟨_, List.reverse_reverse _⟩ 114 | simp [mkBindingList, mkBindingList.core] at * 115 | generalize b.abstractList _ = b 116 | induction xs generalizing b <;> simp_all [mkBindingList.go] 117 | simp [mkBindingList1_congr H.1] 118 | 119 | inductive WF : LocalContext → Prop 120 | | nil : WF ⟨.empty, .empty, .empty⟩ 121 | | cons : 122 | d.fvarId = fv → map.find? fv = none → d.index = arr.size → 123 | WF ⟨map, arr, fvmap⟩ → 124 | WF ⟨map.insert fv d, arr.push d, fvmap⟩ 125 | 126 | theorem WF.map_wf {lctx : LocalContext} : lctx.WF → lctx.fvarIdToDecl.WF 127 | | .nil => .empty 128 | | .cons _ _ _ h2 => .insert h2.map_wf 129 | 130 | theorem WF.decls_wf {lctx : LocalContext} : lctx.WF → lctx.decls.WF 131 | | .nil => .empty 132 | | .cons _ _ _ h2 => .push h2.decls_wf 133 | 134 | attribute [-simp] List.filterMap_reverse in 135 | open scoped _root_.List in 136 | theorem WF.map_toList : WF lctx → 137 | lctx.fvarIdToDecl.toList' ~ lctx.toList.map fun d => (d.fvarId, d) 138 | | .nil => by simp [LocalContext.toList] 139 | | .cons h1 h2 _ h4 => by 140 | subst h1; simp [LocalContext.toList] 141 | refine h4.map_wf.toList'_insert _ _ |>.trans (.cons _ ?_) 142 | rw [List.filter_eq_self.2]; · exact h4.map_toList 143 | simp; rintro _ b h rfl 144 | have := (h4.map_wf.find?_eq _).symm.trans h2 145 | simp [List.lookup_eq_none_iff] at this 146 | exact this _ _ h rfl 147 | 148 | theorem WF.find?_eq_find?_toList (H : WF lctx) : 149 | lctx.find? fv = lctx.toList.find? (fv == ·.fvarId) := by 150 | rw [LocalContext.find?, H.map_wf.find?_eq, 151 | H.map_toList.lookup_eq H.map_wf.nodupKeys, List.map_fst_lookup] 152 | 153 | theorem WF.nodup : WF lctx → (lctx.toList.map (·.fvarId)).Nodup 154 | | .nil => .nil 155 | | .cons h1 h2 h3 h4 => by 156 | have := h4.nodup 157 | have := h4.find?_eq_find?_toList.symm.trans h2 158 | simp_all [toList] 159 | simpa [eq_comm] using this 160 | 161 | protected theorem WF.mkLocalDecl 162 | (h1 : WF lctx) (h2 : lctx.find? fv = none) : WF (lctx.mkLocalDecl fv name ty bi kind) := 163 | .cons rfl h2 rfl h1 164 | 165 | protected theorem WF.mkLetDecl 166 | (h1 : WF lctx) (h2 : lctx.find? fv = none) : WF (lctx.mkLetDecl fv name ty val bi kind) := 167 | .cons rfl h2 rfl h1 168 | 169 | @[simp] theorem mkLocalDecl_toList {lctx : LocalContext} : 170 | (lctx.mkLocalDecl fv name ty bi kind).toList = 171 | .cdecl lctx.decls.size fv name ty bi kind :: lctx.toList := by 172 | simp [mkLocalDecl, toList] 173 | 174 | @[simp] theorem mkLetDecl_toList {lctx : LocalContext} : 175 | (lctx.mkLetDecl fv name ty val bi kind).toList = 176 | .ldecl lctx.decls.size fv name ty val bi kind :: lctx.toList := by 177 | simp [mkLetDecl, toList] 178 | 179 | end Lean.LocalContext 180 | 181 | namespace Lean4Lean 182 | 183 | open Lean 184 | open scoped List 185 | 186 | attribute [-simp] List.filterMap_reverse 187 | 188 | variable (env : VEnv) (Us : List Name) (Δ : VLCtx) in 189 | inductive TrLocalDecl : LocalDecl → VLocalDecl → Prop 190 | | vlam : TrExprS env Us Δ ty ty' → env.IsType Us.length Δ.toCtx ty' → 191 | TrLocalDecl (.cdecl n fv name ty bi kind) (.vlam ty') 192 | | vlet : 193 | TrExprS env Us Δ ty ty' → TrExprS env Us Δ val val' → 194 | env.HasType Us.length Δ.toCtx val' ty' → 195 | TrLocalDecl (.ldecl n fv name ty val bi kind) (.vlet ty' val') 196 | 197 | theorem TrLocalDecl.wf : TrLocalDecl env Us Δ d d' → d'.WF env Us.length Δ.toCtx 198 | | .vlam _ h | .vlet _ _ h => h 199 | 200 | def _root_.Lean.LocalDecl.deps : LocalDecl → List FVarId 201 | | .cdecl (type := t) .. => t.fvarsList 202 | | .ldecl (type := t) (value := v) .. => t.fvarsList ++ v.fvarsList 203 | 204 | theorem TrLocalDecl.deps_wf : TrLocalDecl env Us Δ d d' → d.deps ⊆ Δ.fvars 205 | | .vlam h _ => h.fvarsList 206 | | .vlet h1 h2 _ => by simp [LocalDecl.deps, h1.fvarsList, h2.fvarsList] 207 | 208 | variable (env : VEnv) (Us : List Name) in 209 | inductive TrLCtx' : List LocalDecl → VLCtx → Prop 210 | | nil : TrLCtx' [] [] 211 | | cons : 212 | TrLCtx' ds Δ → TrLocalDecl env Us Δ d d' → 213 | TrLCtx' (d :: ds) ((some (d.fvarId, d.deps), d') :: Δ) 214 | 215 | def TrLCtx (env : VEnv) (Us : List Name) (lctx : LocalContext) (Δ : VLCtx) : Prop := 216 | lctx.WF ∧ TrLCtx' env Us lctx.toList Δ 217 | 218 | theorem TrLCtx.nil {env : VEnv} {Us : List Name} : TrLCtx env Us {} [] := ⟨.nil, .nil⟩ 219 | 220 | theorem TrLCtx'.noBV : TrLCtx' env Us ds Δ → Δ.NoBV 221 | | .nil => rfl 222 | | .cons h _ => h.noBV 223 | 224 | theorem TrLCtx'.forall₂ : 225 | TrLCtx' env Us ds Δ → ds.Forall₂ Δ (R := fun d d' => d'.1 = some (d.fvarId, d.deps)) 226 | | .nil => by simp 227 | | .cons h _ => by simp; exact h.forall₂ 228 | 229 | theorem TrLCtx'.fvars_eq (H : TrLCtx' env Us ds Δ) : ds.map (·.fvarId) = Δ.fvars := by 230 | simp [VLCtx.fvars] 231 | induction H with 232 | | nil => rfl 233 | | cons h1 _ ih => simp [← ih] 234 | 235 | theorem TrLCtx.fvars_eq (H : TrLCtx env Us lctx Δ) : lctx.fvars = Δ.fvars := 236 | H.2.fvars_eq 237 | 238 | theorem TrLCtx'.find?_eq_some (H : TrLCtx' env Us ds Δ) : 239 | (∃ d, ds.find? (fv == ·.fvarId) = some d) ↔ fv ∈ Δ.fvars := by 240 | rw [← Option.isSome_iff_exists, List.find?_isSome] 241 | induction H with simp 242 | | @cons _ _ d d' _ _ ih => simp [← ih] 243 | 244 | theorem TrLCtx'.find?_isSome (H : TrLCtx' env Us ds Δ) : 245 | (ds.find? (fv == ·.fvarId)).isSome = (Δ.find? (.inr fv)).isSome := by 246 | rw [Bool.eq_iff_iff, Option.isSome_iff_exists, Option.isSome_iff_exists, 247 | H.find?_eq_some, VLCtx.find?_eq_some] 248 | 249 | theorem TrLCtx.find?_isSome (H : TrLCtx env Us lctx Δ) : 250 | (lctx.find? fv).isSome = (Δ.find? (.inr fv)).isSome := by 251 | rw [H.1.find?_eq_find?_toList, H.2.find?_isSome] 252 | 253 | theorem TrLCtx.find?_eq_some (H : TrLCtx env Us lctx Δ) : 254 | (∃ d, lctx.find? fv = some d) ↔ fv ∈ Δ.fvars := by 255 | rw [H.1.find?_eq_find?_toList, H.2.find?_eq_some] 256 | 257 | theorem TrLCtx.find?_eq_none (H : TrLCtx env Us lctx Δ) : 258 | lctx.find? fv = none ↔ ¬fv ∈ Δ.fvars := by simp [← H.find?_eq_some] 259 | 260 | theorem TrLCtx.contains (H : TrLCtx env Us lctx Δ) : lctx.contains fv ↔ fv ∈ Δ.fvars := by 261 | rw [LocalContext.contains, PersistentHashMap.find?_isSome, Option.isSome_iff_exists] 262 | exact H.find?_eq_some 263 | 264 | theorem TrLCtx'.wf : TrLCtx' env Us ds Δ → (ds.map (·.fvarId)).Nodup → Δ.WF env Us.length 265 | | .nil, _ => ⟨⟩ 266 | | .cons h1 h2, .cons H1 H2 => by 267 | refine ⟨h1.wf H2, fun _ _ => ?_, h2.wf⟩ 268 | rintro ⟨⟩; exact ⟨by simpa [← h1.find?_eq_some] using H1, h2.deps_wf⟩ 269 | 270 | theorem TrLCtx.wf (H : TrLCtx env Us lctx Δ) : Δ.WF env Us.length := H.2.wf H.1.nodup 271 | 272 | def _root_.Lean.LocalDecl.value' : LocalDecl → Expr 273 | | .ldecl (value := v) .. => v 274 | | .cdecl (fvarId := fv) .. => .fvar fv 275 | 276 | theorem TrLCtx'.find?_of_mem (henv : env.WF) (H : TrLCtx' env Us ds Δ) 277 | (nd : (ds.map (·.fvarId)).Nodup) (hm : decl ∈ ds) : 278 | ∃ e A, Δ.find? (.inr decl.fvarId) = some (e, A) ∧ 279 | FVarsBelow Δ (.fvar decl.fvarId) decl.value' ∧ FVarsBelow Δ (.fvar decl.fvarId) decl.type ∧ 280 | TrExprS env Us Δ decl.value' e ∧ TrExprS env Us Δ decl.type A := by 281 | have := H.wf nd 282 | match H with 283 | | .nil => cases hm 284 | | .cons (ds := ds) h1 h2 => 285 | simp [VLCtx.find?, VLCtx.next] 286 | obtain _ | ⟨_, hm : decl ∈ ds⟩ := hm 287 | · simp [and_assoc] 288 | cases h2 with 289 | | vlam h2 h3 => 290 | refine ⟨.rfl, ?_, .fvar <| by simp [VLCtx.find?, VLCtx.next, LocalDecl.fvarId]; rfl, ?_⟩ 291 | · intro P hP he; exact fvarsIn_iff.2 ⟨hP.2 he, h2.fvarsIn.mono fun _ _ => ⟨⟩⟩ 292 | · exact h2.weakFV henv (.skip_fvar _ _ .refl) this 293 | | vlet h2 h3 => 294 | refine ⟨?_, ?_, ?_, ?_⟩ 295 | · intro P hP he; have := hP.2 he; simp [LocalDecl.deps, or_imp, forall_and] at this 296 | exact fvarsIn_iff.2 ⟨this.2, h3.fvarsIn.mono fun _ _ => ⟨⟩⟩ 297 | · intro P hP he; have := hP.2 he; simp [LocalDecl.deps, or_imp, forall_and] at this 298 | exact fvarsIn_iff.2 ⟨this.1, h2.fvarsIn.mono fun _ _ => ⟨⟩⟩ 299 | · simpa [VLocalDecl.depth] using h3.weakFV henv (.skip_fvar _ _ .refl) this 300 | · simpa [VLocalDecl.depth] using h2.weakFV henv (.skip_fvar _ _ .refl) this 301 | · simp at nd; rw [if_neg (by simpa using Ne.symm (nd.1 _ hm))]; simp 302 | have ⟨_, _, h1, h2, h3, h4, h5⟩ := h1.find?_of_mem henv nd.2 hm 303 | refine ⟨_, _, ⟨_, _, h1, rfl, rfl⟩, fun _ h => h2 _ h.1, fun _ h => h3 _ h.1, ?_, ?_⟩ 304 | · simpa using h4.weakFV henv (.skip_fvar _ _ .refl) this 305 | · simpa using h5.weakFV henv (.skip_fvar _ _ .refl) this 306 | 307 | theorem TrLCtx.find?_of_mem (henv : env.WF) (H : TrLCtx env Us lctx Δ) 308 | (hm : decl ∈ lctx.toList) : 309 | ∃ e A, Δ.find? (.inr decl.fvarId) = some (e, A) ∧ 310 | FVarsBelow Δ (.fvar decl.fvarId) decl.value' ∧ FVarsBelow Δ (.fvar decl.fvarId) decl.type ∧ 311 | TrExprS env Us Δ decl.value' e ∧ TrExprS env Us Δ decl.type A := 312 | H.2.find?_of_mem henv H.1.nodup hm 313 | 314 | theorem TrLCtx.mkLocalDecl 315 | (h1 : TrLCtx env Us lctx Δ) (h2 : lctx.find? fv = none) (h3 : TrExprS env Us Δ ty ty') 316 | (h4 : env.IsType Us.length Δ.toCtx ty') : 317 | TrLCtx env Us (lctx.mkLocalDecl fv name ty bi kind) 318 | ((some (fv, ty.fvarsList), .vlam ty') :: Δ) := 319 | ⟨h1.1.mkLocalDecl h2, by simpa using .cons h1.2 (.vlam h3 h4)⟩ 320 | 321 | theorem TrLCtx.mkLetDecl 322 | (h1 : TrLCtx env Us lctx Δ) (h2 : lctx.find? fv = none) 323 | (h3 : TrExprS env Us Δ ty ty') (h4 : TrExprS env Us Δ val val') 324 | (h5 : env.HasType Us.length Δ.toCtx val' ty') : 325 | TrLCtx env Us (lctx.mkLetDecl fv name ty val bi kind) 326 | ((some (fv, ty.fvarsList ++ val.fvarsList), .vlet ty' val') :: Δ) := 327 | ⟨h1.1.mkLetDecl h2, by simpa using .cons h1.2 (.vlet h3 h4 h5)⟩ 328 | --------------------------------------------------------------------------------