├── .gitignore ├── Lean4Lean.lean ├── Lean4Lean ├── Declaration.lean ├── Environment.lean ├── Environment │ └── Basic.lean ├── EquivManager.lean ├── Expr.lean ├── ForEachExprV.lean ├── Inductive │ ├── Add.lean │ └── Reduce.lean ├── Instantiate.lean ├── Level.lean ├── List.lean ├── LocalContext.lean ├── Primitive.lean ├── Quot.lean ├── Std │ ├── Basic.lean │ ├── Control.lean │ ├── HashMap.lean │ ├── NodupKeys.lean │ ├── PersistentHashMap.lean │ ├── SMap.lean │ ├── ToExpr.lean │ └── Variable!.lean ├── Stream.lean ├── Theory.lean ├── Theory │ ├── Inductive.lean │ ├── Meta.lean │ ├── Quot.lean │ ├── Typing │ │ ├── Basic.lean │ │ ├── Env.lean │ │ ├── EnvLemmas.lean │ │ ├── InductiveLemmas.lean │ │ ├── Lemmas.lean │ │ ├── Meta.lean │ │ ├── NormalEq.lean │ │ ├── ParallelReduction.lean │ │ ├── QuotLemmas.lean │ │ ├── Stratified.lean │ │ ├── StratifiedUntyped.lean │ │ ├── Strong.lean │ │ └── UniqueTyping.lean │ ├── VDecl.lean │ ├── VEnv.lean │ ├── VExpr.lean │ └── VLevel.lean ├── TypeChecker.lean ├── Verify.lean └── Verify │ ├── Axioms.lean │ ├── Environment.lean │ ├── Expr.lean │ ├── Level.lean │ ├── LocalContext.lean │ ├── NameGenerator.lean │ ├── TypeChecker │ └── Theorems.lean │ ├── Typing │ ├── ConditionallyTyped.lean │ ├── EnvLemmas.lean │ ├── Expr.lean │ └── Lemmas.lean │ └── VLCtx.lean ├── Main.lean ├── README.md ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /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/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 | -------------------------------------------------------------------------------- /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 ← check v.type v.levelParams 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 := {}) do 37 | let valType ← TypeChecker.check v.value v.levelParams 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 := {}) do 44 | checkConstantVal env v.toConstantVal (← checkPrimitiveDef env v) 45 | checkNoMVarNoFVar env v.name v.value 46 | let valType ← TypeChecker.check v.value v.levelParams 47 | if !(← isDefEq valType v.type) then 48 | throw <| .declTypeMismatch env (.defnDecl v) valType 49 | return env.add (.defnInfo v) 50 | 51 | def addTheorem (env : Environment) (v : TheoremVal) (check := true) : 52 | Except Exception Environment := do 53 | if check then 54 | -- TODO(Leo): we must add support for handling tasks here 55 | M.run env (safety := .safe) (lctx := {}) do 56 | if !(← isProp v.type) then 57 | throw <| .thmTypeIsNotProp env v.name v.type 58 | checkConstantVal env v.toConstantVal 59 | checkNoMVarNoFVar env v.name v.value 60 | let valType ← TypeChecker.check v.value v.levelParams 61 | if !(← isDefEq valType v.type) then 62 | throw <| .declTypeMismatch env (.thmDecl v) valType 63 | return env.add (.thmInfo v) 64 | 65 | def addOpaque (env : Environment) (v : OpaqueVal) (check := true) : 66 | Except Exception Environment := do 67 | if check then 68 | M.run env (safety := .safe) (lctx := {}) do 69 | checkConstantVal env v.toConstantVal 70 | let valType ← TypeChecker.check v.value v.levelParams 71 | if !(← isDefEq valType v.type) then 72 | throw <| .declTypeMismatch env (.opaqueDecl v) valType 73 | return env.add (.opaqueInfo v) 74 | 75 | def addMutual (env : Environment) (vs : List DefinitionVal) (check := true) : 76 | Except Exception Environment := do 77 | let v₀ :: _ := vs | throw <| .other "invalid empty mutual definition" 78 | if let .safe := v₀.safety then 79 | throw <| .other "invalid mutual definition, declaration is not tagged as unsafe/partial" 80 | if check then 81 | M.run env (safety := v₀.safety) (lctx := {}) do 82 | for v in vs do 83 | if v.safety != v₀.safety then 84 | throw <| .other 85 | "invalid mutual definition, declarations must have the same safety annotation" 86 | checkConstantVal env v.toConstantVal 87 | let mut env' := env 88 | for v in vs do 89 | env' := env'.add (.defnInfo v) 90 | if check then 91 | M.run env' (safety := v₀.safety) (lctx := {}) do 92 | for v in vs do 93 | checkNoMVarNoFVar env' v.name v.value 94 | let valType ← TypeChecker.check v.value v.levelParams 95 | if !(← isDefEq valType v.type) then 96 | throw <| .declTypeMismatch env' (.mutualDefnDecl vs) valType 97 | return env' 98 | 99 | /-- Type check given declaration and add it to the environment -/ 100 | def addDecl (env : Environment) (decl : Declaration) (check := true) : 101 | Except Exception Environment := do 102 | match decl with 103 | | .axiomDecl v => addAxiom env v check 104 | | .defnDecl v => addDefinition env v check 105 | | .thmDecl v => addTheorem env v check 106 | | .opaqueDecl v => addOpaque env v check 107 | | .mutualDefnDecl v => addMutual env v check 108 | | .quotDecl => addQuot env 109 | | .inductDecl lparams nparams types isUnsafe => 110 | let allowPrimitive ← checkPrimitiveInductive env lparams nparams types isUnsafe 111 | addInductive env lparams nparams types isUnsafe allowPrimitive 112 | -------------------------------------------------------------------------------- /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 | ``String, ``String.mk] 40 | 41 | /-- 42 | Returns true iff `constName` is a non-recursive inductive datatype that has only one constructor and no indices. 43 | 44 | Such types have special kernel support. This must be in sync with `is_structure_like`. 45 | -/ 46 | def isStructureLike (env : Environment) (constName : Name) : Bool := 47 | match env.find? constName with 48 | | some (.inductInfo { isRec := false, ctors := [_], numIndices := 0, .. }) => true 49 | | _ => false 50 | 51 | def checkName (env : Environment) (n : Name) 52 | (allowPrimitive := false) : Except Exception Unit := do 53 | if env.contains n then 54 | throw <| .alreadyDeclared env n 55 | unless allowPrimitive do 56 | if primitives.contains n then 57 | throw <| .other s!"unexpected use of primitive name {n}" 58 | 59 | open private subsumesInfo Kernel.Environment.mk moduleNames moduleNameMap parts 60 | from Lean.Environment 61 | 62 | def empty (mainModule : Name) (trustLevel : UInt32 := 0) : Environment := 63 | Kernel.Environment.mk 64 | (const2ModIdx := {}) 65 | (constants := {}) 66 | (quotInit := false) 67 | (diagnostics := {}) 68 | (extraConstNames := {}) 69 | (header := { mainModule, trustLevel }) 70 | (extensions := #[]) 71 | 72 | def throwAlreadyImported (s : ImportState) (const2ModIdx : Std.HashMap Name ModuleIdx) 73 | (modIdx : Nat) (cname : Name) : Except Exception α := do 74 | let modName := (moduleNames s)[modIdx]! 75 | let constModName := (moduleNames s)[const2ModIdx[cname]!.toNat]! 76 | throw <| .other 77 | s!"import {modName} failed, environment already contains '{cname}' from {constModName}" 78 | 79 | def finalizeImport (s : ImportState) (imports : Array Import) (mainModule : Name) 80 | (trustLevel : UInt32 := 0) : Except Exception Environment := do 81 | let modules := (moduleNames s).filterMap ((moduleNameMap s)[·]?) 82 | let moduleData ← modules.mapM fun mod => do 83 | let some data := mod.mainModule? | 84 | throw <| .other s!"missing data file for module {mod.module}" 85 | return data 86 | let numPrivateConsts := moduleData.foldl (init := 0) fun numPrivateConsts data => Id.run do 87 | numPrivateConsts + data.constants.size + data.extraConstNames.size 88 | let mut const2ModIdx := .emptyWithCapacity (capacity := numPrivateConsts) 89 | let mut privateConstantMap := .emptyWithCapacity (capacity := numPrivateConsts) 90 | for h : modIdx in [0:moduleData.size] do 91 | let data := moduleData[modIdx] 92 | for cname in data.constNames, cinfo in data.constants do 93 | match privateConstantMap.getThenInsertIfNew? cname cinfo with 94 | | (cinfoPrev?, constantMap') => 95 | privateConstantMap := constantMap' 96 | if let some cinfoPrev := cinfoPrev? then 97 | -- Recall that the map has not been modified when `cinfoPrev? = some _`. 98 | if subsumesInfo cinfo cinfoPrev then 99 | privateConstantMap := privateConstantMap.insert cname cinfo 100 | else if !subsumesInfo cinfoPrev cinfo then 101 | throwAlreadyImported s const2ModIdx modIdx cname 102 | const2ModIdx := const2ModIdx.insertIfNew cname modIdx 103 | for cname in data.extraConstNames do 104 | const2ModIdx := const2ModIdx.insertIfNew cname modIdx 105 | 106 | return Kernel.Environment.mk 107 | (const2ModIdx := const2ModIdx) 108 | (constants := SMap.fromHashMap privateConstantMap false) 109 | (quotInit := !imports.isEmpty) -- We assume `Init.Prelude` initializes quotient module 110 | (diagnostics := {}) 111 | (extraConstNames := {}) 112 | (extensions := #[]) 113 | (header := { 114 | trustLevel, imports, mainModule 115 | regions := modules.flatMap (parts · |>.map (·.2)) 116 | moduleNames := moduleNames s 117 | moduleData 118 | }) 119 | -------------------------------------------------------------------------------- /Lean4Lean/EquivManager.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Data.UnionFind.Basic 2 | 3 | namespace Lean4Lean 4 | open Lean 5 | 6 | abbrev EquivManager.NodeRef := Nat 7 | open EquivManager 8 | 9 | structure EquivManager where 10 | uf : Batteries.UnionFind := {} 11 | toNodeMap : ExprMap NodeRef := {} 12 | 13 | namespace EquivManager 14 | 15 | def find (n : NodeRef) : StateM EquivManager NodeRef := fun m => 16 | if h : n < m.uf.size then 17 | let ⟨uf, root, _⟩ := m.uf.find ⟨n, h⟩ 18 | (root, { m with uf }) 19 | else (n, m) 20 | 21 | def merge (m : EquivManager) (n1 n2 : NodeRef) : EquivManager := 22 | if h1 : n1 < m.uf.size then 23 | if h2 : n2 < m.uf.size then 24 | { m with uf := m.uf.union ⟨n1, h1⟩ ⟨n2, h2⟩ } 25 | else m 26 | else m 27 | 28 | def toNode (e : Expr) : StateM EquivManager NodeRef := fun m => do 29 | if let some r := m.toNodeMap[e]? then 30 | return (r, m) 31 | let { uf, toNodeMap } := m 32 | let r := uf.size 33 | (r, { uf := uf.push, toNodeMap := toNodeMap.insert e r }) 34 | 35 | variable (useHash : Bool) in 36 | def isEquiv (e1 e2 : Expr) : StateM EquivManager Bool := do 37 | -- FIXME: find a way to do this with withPtrEq or benchmark how important it is 38 | if unsafe ptrEq 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 | 60 | def addEquiv (m : EquivManager) (e1 e2 : Expr) : EquivManager := 61 | let (r1, m) := toNode e1 m 62 | let (r2, m) := toNode e2 m 63 | merge m r1 r2 64 | -------------------------------------------------------------------------------- /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 | namespace ReplaceImpl 11 | 12 | unsafe abbrev ReplaceT := StateT (PtrMap Expr Expr) 13 | 14 | @[inline] 15 | unsafe def cacheT [Monad m] (key : Expr) (result : Expr) : ReplaceT m Expr := do 16 | modify (·.insert key result) 17 | pure result 18 | 19 | @[specialize] 20 | unsafe def replaceUnsafeT [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : ReplaceT m Expr := do 21 | let rec @[specialize] visit (e : Expr) := do 22 | if let some result := (← get).find? e then 23 | return result 24 | match ← f? e with 25 | | some eNew => cacheT e eNew 26 | | none => match e with 27 | | Expr.forallE _ d b _ => cacheT e <| e.updateForallE! (← visit d) (← visit b) 28 | | Expr.lam _ d b _ => cacheT e <| e.updateLambdaE! (← visit d) (← visit b) 29 | | Expr.mdata _ b => cacheT e <| e.updateMData! (← visit b) 30 | | Expr.letE _ t v b _ => cacheT e <| e.updateLet! (← visit t) (← visit v) (← visit b) 31 | | Expr.app f a => cacheT e <| e.updateApp! (← visit f) (← visit a) 32 | | Expr.proj _ _ b => cacheT e <| e.updateProj! (← visit b) 33 | | e => pure e 34 | visit e 35 | 36 | @[inline] 37 | unsafe def replaceUnsafe' [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := 38 | (replaceUnsafeT f? e).run' mkPtrMap 39 | 40 | end ReplaceImpl 41 | 42 | /- TODO: use withPtrAddr, withPtrEq to avoid unsafe tricks above. 43 | We also need an invariant at `State` and proofs for the `uget` operations. -/ 44 | 45 | @[specialize] 46 | def replaceNoCacheT [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := do 47 | match ← f? e with 48 | | some eNew => pure eNew 49 | | none => match e with 50 | | .forallE _ d b _ => 51 | return e.updateForallE! (← replaceNoCacheT f? d) (← replaceNoCacheT f? b) 52 | | .lam _ d b _ => 53 | return e.updateLambdaE! (← replaceNoCacheT f? d) (← replaceNoCacheT f? b) 54 | | .mdata _ b => 55 | return e.updateMData! (← replaceNoCacheT f? b) 56 | | .letE _ t v b _ => 57 | return e.updateLet! (← replaceNoCacheT f? t) (← replaceNoCacheT f? v) (← replaceNoCacheT f? b) 58 | | .app f a => 59 | return e.updateApp! (← replaceNoCacheT f? f) (← replaceNoCacheT f? a) 60 | | .proj _ _ b => 61 | return e.updateProj! (← replaceNoCacheT f? b) 62 | | e => return e 63 | 64 | @[implemented_by ReplaceImpl.replaceUnsafe'] 65 | partial def replaceM [Monad m] (f? : Expr → m (Option Expr)) (e : Expr) : m Expr := 66 | e.replaceNoCacheT f? 67 | 68 | def natZero : Expr := .const ``Nat.zero [] 69 | def natSucc : Expr := .const ``Nat.succ [] 70 | 71 | def isConstructorApp?' (env : Kernel.Environment) (e : Expr) : Option Name := do 72 | let .const fn _ := e.getAppFn | none 73 | let .ctorInfo _ ← env.find? fn | none 74 | return fn 75 | 76 | def natLitToConstructor : Nat → Expr 77 | | 0 => natZero 78 | | n+1 => .app natSucc (.lit (.natVal n)) 79 | 80 | def strLitToConstructor (s : String) : Expr := 81 | let char := .const ``Char [] 82 | let listNil := .app (.const ``List.nil [.zero]) char 83 | let listCons := .app (.const ``List.cons [.zero]) char 84 | let stringMk := .const ``String.mk [] 85 | let charOfNat := .const ``Char.ofNat [] 86 | .app stringMk <| s.foldr (init := listNil) fun c e => 87 | .app (.app listCons <| .app charOfNat (.lit (.natVal c.toNat))) e 88 | 89 | end Expr 90 | 91 | def Literal.toConstructor : Literal → Expr 92 | | .natVal n => .natLitToConstructor n 93 | | .strVal s => .strLitToConstructor s 94 | -------------------------------------------------------------------------------- /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/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/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 : Id Expr := 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 | loop 0 fn 28 | 29 | end Expr 30 | -------------------------------------------------------------------------------- /Lean4Lean/Level.lean: -------------------------------------------------------------------------------- 1 | import Lean.Level 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 | def isEquivList : List Level → List Level → Bool := List.all2 isEquiv 24 | 25 | def geq' (u v : Level) : Bool := -- pending lean4#2689 26 | go u.normalize v.normalize 27 | where 28 | go (u v : Level) : Bool := 29 | u == v || 30 | let k := fun () => 31 | match v with 32 | | imax v₁ v₂ => go u v₁ && go u v₂ 33 | | _ => 34 | let v' := v.getLevelOffset 35 | (u.getLevelOffset == v' || v'.isZero) 36 | && u.getOffset ≥ v.getOffset 37 | match u, v with 38 | | _, zero => true 39 | | u, max v₁ v₂ => go u v₁ && go u v₂ 40 | | max u₁ u₂, v => go u₁ v || go u₂ v || k () 41 | | imax _ u₂, v => go u₂ v 42 | | succ u, succ v => go u v 43 | | _, _ => k () 44 | termination_by (u, v) 45 | -------------------------------------------------------------------------------- /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/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/Primitive.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.TypeChecker 2 | import Lean4Lean.Environment.Basic 3 | 4 | namespace Lean4Lean 5 | namespace Environment 6 | open Lean hiding Environment Exception 7 | open Kernel TypeChecker 8 | 9 | open private Lean.Kernel.Environment.add from Lean.Environment 10 | 11 | def checkPrimitiveDef (env : Environment) (v : DefinitionVal) : M Bool := do 12 | let fail {α} : M α := throw <| .other s!"invalid form for primitive def {v.name}" 13 | let nat := .const ``Nat [] 14 | let bool := .const ``Bool [] 15 | let tru := .const ``Bool.true [] 16 | let fal := .const ``Bool.false [] 17 | let zero := .const ``Nat.zero [] 18 | let succ := mkApp (.const ``Nat.succ []) 19 | let pred := mkApp (.const ``Nat.pred []) 20 | let add := mkApp2 (.const ``Nat.add []) 21 | let mul := mkApp2 (.const ``Nat.mul []) 22 | let mod := mkApp2 (.const ``Nat.mod []) 23 | let defeq1 a b := TypeChecker.isDefEq (.arrow nat a) (.arrow nat b) 24 | let defeq2 a b := defeq1 (.arrow nat a) (.arrow nat b) 25 | let x := .bvar 0 26 | let y := .bvar 1 27 | match v.name with 28 | | ``Nat.add => 29 | unless env.contains ``Nat && v.levelParams.isEmpty do fail 30 | -- add : Nat → Nat → Nat 31 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 32 | let add := mkApp2 v.value 33 | -- add x 0 ≡ x 34 | unless ← defeq1 (add x zero) x do fail 35 | -- add y (succ x) ≡ succ (add y x) 36 | unless ← defeq2 (add y (succ x)) (succ (add y x)) do fail 37 | | ``Nat.pred => 38 | unless env.contains ``Nat && v.levelParams.isEmpty do fail 39 | -- pred : Nat → Nat 40 | unless ← TypeChecker.isDefEq v.type (.arrow nat nat) do fail 41 | let pred := mkApp v.value 42 | unless ← TypeChecker.isDefEq (pred zero) zero do fail 43 | unless ← defeq1 (pred (succ x)) x do fail 44 | | ``Nat.sub => 45 | unless env.contains ``Nat.pred && v.levelParams.isEmpty do fail 46 | -- sub : Nat → Nat → Nat 47 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 48 | let sub := mkApp2 v.value 49 | unless ← defeq1 (sub x zero) x do fail 50 | unless ← defeq2 (sub y (succ x)) (pred (sub y x)) do fail 51 | | ``Nat.mul => 52 | unless env.contains ``Nat.add && v.levelParams.isEmpty do fail 53 | -- mul : Nat → Nat → Nat 54 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 55 | let mul := mkApp2 v.value 56 | unless ← defeq1 (mul x zero) zero do fail 57 | unless ← defeq2 (mul y (succ x)) (add (mul y x) y) do fail 58 | | ``Nat.pow => 59 | unless env.contains ``Nat.mul && v.levelParams.isEmpty do fail 60 | -- pow : Nat → Nat → Nat 61 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 62 | let pow := mkApp2 v.value 63 | unless ← defeq1 (pow x zero) (succ zero) do fail 64 | unless ← defeq2 (pow y (succ x)) (mul (pow y x) y) do fail 65 | | ``Nat.mod => 66 | unless env.contains ``Nat.sub && v.levelParams.isEmpty do fail 67 | -- mod : Nat → Nat → Nat 68 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 69 | let mod := mkApp2 v.value 70 | unless ← defeq1 (mod zero x) zero do fail 71 | return true -- TODO 72 | | ``Nat.div => 73 | unless env.contains ``Nat.sub && v.levelParams.isEmpty do fail 74 | -- div : Nat → Nat → Nat 75 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 76 | return true -- TODO 77 | | ``Nat.gcd => 78 | unless env.contains ``Nat.mod && v.levelParams.isEmpty do fail 79 | -- gcd : Nat → Nat → Nat 80 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat nat)) do fail 81 | let gcd := mkApp2 v.value 82 | unless ← defeq1 (gcd zero x) x do fail 83 | unless ← defeq2 (gcd (succ y) x) (gcd (mod x (succ y)) (succ y)) do fail 84 | | ``Nat.beq => 85 | unless env.contains ``Nat && env.contains ``Bool && v.levelParams.isEmpty do fail 86 | -- beq : Nat → Nat → Bool 87 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat bool)) do fail 88 | let beq := mkApp2 v.value 89 | unless ← TypeChecker.isDefEq (beq zero zero) tru do fail 90 | unless ← defeq1 (beq zero (succ x)) fal do fail 91 | unless ← defeq1 (beq (succ x) zero) fal do fail 92 | unless ← defeq2 (beq (succ y) (succ x)) (beq y x) do fail 93 | | ``Nat.ble => 94 | unless env.contains ``Nat && env.contains ``Bool && v.levelParams.isEmpty do fail 95 | -- ble : Nat → Nat → Bool 96 | unless ← TypeChecker.isDefEq v.type (.arrow nat (.arrow nat bool)) do fail 97 | let ble := mkApp2 v.value 98 | unless ← TypeChecker.isDefEq (ble zero zero) tru do fail 99 | unless ← defeq1 (ble zero (succ x)) tru do fail 100 | unless ← defeq1 (ble (succ x) zero) fal do fail 101 | unless ← defeq2 (ble (succ y) (succ x)) (ble y x) do fail 102 | | _ => return false 103 | return true 104 | 105 | def checkPrimitiveInductive (env : Environment) (lparams : List Name) (nparams : Nat) 106 | (types : List InductiveType) (isUnsafe : Bool) : Except Exception Bool := do 107 | unless !isUnsafe && lparams.isEmpty && nparams == 0 do return false 108 | let [type] := types | return false 109 | unless type.type == .sort (.succ .zero) do return false 110 | let fail {α} : Except Exception α := 111 | throw <| .other s!"invalid form for primitive inductive {type.name}" 112 | match type.name with 113 | | ``Bool => 114 | let [⟨``Bool.false, .const ``Bool []⟩, ⟨``Bool.true, .const ``Bool []⟩] := type.ctors | fail 115 | | ``Nat => 116 | let [ 117 | ⟨``Nat.zero, .const ``Nat []⟩, 118 | ⟨``Nat.succ, .forallE _ (.const ``Nat []) (.const ``Nat []) _⟩ 119 | ] := type.ctors | fail 120 | | ``String => 121 | let [⟨``String.mk, 122 | .forallE _ (.app (.const ``List [.zero]) (.const ``Char [])) (.const ``String []) _ 123 | ⟩] := type.ctors | fail 124 | M.run env (safety := .safe) (lctx := {}) do 125 | -- We need the following definitions for `strLitToConstructor` to work: 126 | -- Nat : Type (this is primitive so checking for existence suffices) 127 | let nat := .const ``Nat [] 128 | unless env.contains ``Nat do fail 129 | -- Char : Type 130 | let char := .const ``Char [] 131 | _ ← TypeChecker.ensureType char 132 | -- List Char : Type 133 | let listchar := mkApp (.const ``List [.zero]) char 134 | _ ← TypeChecker.ensureType listchar 135 | -- @List.nil.{0} Char : List Char 136 | let listNil := .app (.const ``List.nil [.zero]) char 137 | unless ← TypeChecker.isDefEq (← TypeChecker.check listNil []) listchar do fail 138 | -- @List.cons.{0} Char : List Char 139 | let listCons := .app (.const ``List.cons [.zero]) char 140 | unless ← TypeChecker.isDefEq (← TypeChecker.check listCons []) 141 | (.arrow char (.arrow listchar listchar)) do fail 142 | -- String.mk : List Char → String (already checked) 143 | -- @Char.ofNat : Nat → Char 144 | let charOfNat := .const ``Char.ofNat [] 145 | unless ← TypeChecker.isDefEq (← TypeChecker.check charOfNat []) (.arrow nat char) do fail 146 | | _ => return false 147 | return true 148 | -------------------------------------------------------------------------------- /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/Std/Basic.lean: -------------------------------------------------------------------------------- 1 | import Batteries.CodeAction 2 | import Batteries.Data.Array.Lemmas 3 | import Batteries.Data.HashMap.Basic 4 | 5 | attribute [simp] Option.bind_eq_some List.filterMap_cons 6 | 7 | protected theorem Nat.le_iff_exists_add {a b : Nat} : a ≤ b ↔ ∃ c, b = a + c := 8 | ⟨fun h => ⟨_, (Nat.add_sub_cancel' h).symm⟩, fun ⟨_, h⟩ => h ▸ Nat.le_add_right ..⟩ 9 | 10 | protected theorem Nat.le_iff_exists_add' {a b : Nat} : a ≤ b ↔ ∃ c, b = c + a := by 11 | simp [Nat.add_comm, Nat.le_iff_exists_add] 12 | 13 | protected theorem List.Forall₂.rfl 14 | {R : α → α → Prop} {l : List α} (h : ∀ a ∈ l, R a a) : l.Forall₂ R l := by 15 | induction l with 16 | | nil => constructor 17 | | cons _ _ ih => simp at h; exact .cons h.1 (ih h.2) 18 | 19 | @[simp] theorem List.forall₂_nil_left_iff {l} : Forall₂ R nil l ↔ l = nil := 20 | ⟨fun H => by cases H; rfl, by rintro rfl; exact Forall₂.nil⟩ 21 | 22 | @[simp] theorem List.forall₂_nil_right_iff {l} : Forall₂ R l nil ↔ l = nil := 23 | ⟨fun H => by cases H; rfl, by rintro rfl; exact Forall₂.nil⟩ 24 | 25 | theorem List.forall₂_cons_left_iff {a l u} : 26 | Forall₂ R (a :: l) u ↔ ∃ b u', R a b ∧ Forall₂ R l u' ∧ u = b :: u' := 27 | Iff.intro 28 | (fun h => match u, h with | b :: u', Forall₂.cons h₁ h₂ => ⟨b, u', h₁, h₂, rfl⟩) 29 | (fun h => match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂) 30 | 31 | theorem List.forall₂_cons_right_iff {b l u} : 32 | Forall₂ R u (b :: l) ↔ ∃ a u', R a b ∧ Forall₂ R u' l ∧ u = a :: u' := 33 | Iff.intro 34 | (fun h => match u, h with | b :: u', Forall₂.cons h₁ h₂ => ⟨b, u', h₁, h₂, rfl⟩) 35 | (fun h => match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂) 36 | 37 | theorem List.Forall₂.imp (H : ∀ a b, R a b → S a b) 38 | {l₁ l₂} (h : Forall₂ R l₁ l₂) : Forall₂ S l₁ l₂ := by 39 | induction h <;> constructor <;> [(apply H; assumption); assumption] 40 | 41 | theorem List.Forall₂.trans (H : ∀ a b c, R a b → S b c → T a c) 42 | {l₁ l₂ l₃} (h₁ : Forall₂ R l₁ l₂) (h₂ : Forall₂ S l₂ l₃) : Forall₂ T l₁ l₃ := by 43 | induction h₁ generalizing l₃ <;> cases h₂ <;> constructor <;> solve_by_elim 44 | 45 | @[simp] theorem List.forall₂_map_left_iff {f : γ → α} : 46 | ∀ {l u}, Forall₂ R (map f l) u ↔ Forall₂ (fun c b => R (f c) b) l u 47 | | [], _ => by simp only [map, forall₂_nil_left_iff] 48 | | a :: l, _ => by simp only [map, forall₂_cons_left_iff, forall₂_map_left_iff] 49 | 50 | @[simp] theorem List.forall₂_map_right_iff {f : γ → β} : 51 | ∀ {l u}, Forall₂ R l (map f u) ↔ Forall₂ (fun a c => R a (f c)) l u 52 | | _, [] => by simp only [map, forall₂_nil_right_iff] 53 | | _, b :: u => by simp only [map, forall₂_cons_right_iff, forall₂_map_right_iff] 54 | 55 | theorem List.Forall₂.flip : ∀ {a b}, Forall₂ (flip R) b a → Forall₂ R a b 56 | | _, _, Forall₂.nil => Forall₂.nil 57 | | _ :: _, _ :: _, Forall₂.cons h₁ h₂ => Forall₂.cons h₁ h₂.flip 58 | 59 | theorem List.Forall₂.forall_exists_l {l₁ l₂} (h : Forall₂ R l₁ l₂) : ∀ a ∈ l₁, ∃ b ∈ l₂, R a b := by 60 | induction h with simp [*] | cons _ _ ih => exact fun a h => .inr (ih _ h) 61 | 62 | theorem List.Forall₂.forall_exists_r {l₁ l₂} (h : Forall₂ R l₁ l₂) : ∀ b ∈ l₂, ∃ a ∈ l₁, R a b := 63 | h.flip.forall_exists_l 64 | 65 | theorem List.Forall₂.length_eq : ∀ {l₁ l₂}, Forall₂ R l₁ l₂ → length l₁ = length l₂ 66 | | _, _, Forall₂.nil => rfl 67 | | _, _, Forall₂.cons _ h₂ => congrArg Nat.succ (Forall₂.length_eq h₂) 68 | 69 | theorem List.map_id''' {f : α → α} (l : List α) (h : ∀ x ∈ l, f x = x) : map f l = l := by 70 | induction l <;> simp_all 71 | 72 | theorem List.map_fst_lookup {f : α → β} [BEq β] (l : List α) (b : β) : 73 | (l.map (fun a => (f a, a))).lookup b = l.find? fun a => b == f a := by 74 | induction l <;> simp_all [lookup, find?] 75 | 76 | def List.All (P : α → Prop) : List α → Prop 77 | | [] => True 78 | | a::as => P a ∧ as.All P 79 | 80 | theorem List.All.imp {P Q : α → Prop} (h : ∀ a, P a → Q a) : ∀ {l : List α}, l.All P → l.All Q 81 | | [] => id 82 | | _::_ => And.imp (h _) (List.All.imp h) 83 | 84 | theorem List.append_eq_append_of_length_le {a b c d : List α} (h : length a ≤ length c) : 85 | a ++ b = c ++ d ↔ ∃ a', c = a ++ a' ∧ b = a' ++ d := by 86 | rw [append_eq_append_iff, or_iff_left_iff_imp] 87 | rintro ⟨c', rfl, rfl⟩ 88 | rw [← Nat.add_zero c.length, length_append, 89 | Nat.add_le_add_iff_left, Nat.le_zero, length_eq_zero_iff] at h 90 | subst h; exact ⟨[], by simp⟩ 91 | 92 | @[simp] theorem List.nodup_reverse {l : List α} : Nodup (reverse l) ↔ Nodup l := 93 | pairwise_reverse.trans <| by simp only [Nodup, Ne, eq_comm] 94 | 95 | theorem List.foldl_congr 96 | (H : ∀ a, ∀ x ∈ l, f a x = g a x) : foldl f a l = foldl g a l := by 97 | induction l generalizing a <;> simp_all 98 | 99 | theorem List.idxOf_eq_length_iff [BEq α] [LawfulBEq α] 100 | {a : α} {l : List α} : idxOf a l = length l ↔ a ∉ l := by 101 | induction l with 102 | | nil => exact iff_of_true rfl not_mem_nil 103 | | cons b l ih => 104 | simp only [length, mem_cons, idxOf_cons, eq_comm] 105 | rw [cond_eq_if] 106 | split <;> rename_i h <;> simp at h 107 | · exact iff_of_false (by rintro ⟨⟩) fun H => H <| Or.inl h.symm 108 | · simp only [Ne.symm h, false_or] 109 | rw [← ih] 110 | exact Nat.succ_inj 111 | 112 | theorem List.idxOf_le_length [BEq α] [LawfulBEq α] 113 | {a : α} {l : List α} : idxOf a l ≤ length l := by 114 | induction l with | nil => exact Nat.le_refl _ | cons b l ih => ?_ 115 | simp only [length, idxOf_cons, cond_eq_if, beq_iff_eq] 116 | by_cases h : b == a 117 | · rw [if_pos h]; exact Nat.zero_le _ 118 | · rw [if_neg h]; exact Nat.succ_le_succ ih 119 | 120 | theorem List.idxOf_lt_length_iff [BEq α] [LawfulBEq α] 121 | {a} {l : List α} : idxOf a l < length l ↔ a ∈ l := 122 | ⟨fun h => Decidable.byContradiction fun al => Nat.ne_of_lt h <| idxOf_eq_length_iff.2 al, 123 | fun al => (Nat.lt_of_le_of_ne idxOf_le_length) fun h => idxOf_eq_length_iff.1 h al⟩ 124 | 125 | instance [BEq α] [LawfulBEq α] : PartialEquivBEq α where 126 | symm h := by simp at *; exact h.symm 127 | trans h1 h2 := by simp at *; exact h1.trans h2 128 | 129 | instance : LawfulBEq Lean.FVarId where 130 | eq_of_beq := @fun ⟨a⟩ ⟨b⟩ h => by cases LawfulBEq.eq_of_beq (α := Lean.Name) h; rfl 131 | rfl := BEq.rfl (α := Lean.Name) 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 | -------------------------------------------------------------------------------- /Lean4Lean/Std/Control.lean: -------------------------------------------------------------------------------- 1 | 2 | instance [Alternative m] : MonadLift Option m := ⟨fun | none => failure | some a => pure a⟩ 3 | -------------------------------------------------------------------------------- /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/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 [eq] 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; simp [h1] at h2; simp [h1, h2] at eq 43 | -------------------------------------------------------------------------------- /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 [Option.orElse, 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 [Option.orElse] 73 | cases wf.disjoint (by simp [PersistentHashMap.find?_isSome, e2]) (Std.HashMap.mem_of_getElem? e1) 74 | -------------------------------------------------------------------------------- /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/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/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.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.EnvLemmas 2 | import Lean4Lean.Theory.Typing.Strong 3 | import Lean4Lean.Theory.Typing.Stratified 4 | import Lean4Lean.Theory.Typing.StratifiedUntyped 5 | import Lean4Lean.Theory.Typing.ParallelReduction 6 | import Lean4Lean.Theory.Typing.UniqueTyping 7 | -------------------------------------------------------------------------------- /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 | -------------------------------------------------------------------------------- /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.find? 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/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/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 | -------------------------------------------------------------------------------- /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/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/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/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 IsDefEq.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 IsDefEq.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 IsDefEq.sort ?_; decide 42 | | refine IsDefEq.bvar ?_; lookup_tac 43 | | refine IsDefEq.app' (A := ?_) (B := ?_) ?_ ?_ ?_ <;> [skip; skip; type_tac; type_tac; exact rfl] 44 | | refine IsDefEq.const' (ci := ?_) ?_ ?_ ?_ ?_ <;> [skip; assumption; decide; decide; exact rfl] 45 | | refine HasType.lam (u := ?_) ?_ ?_ <;> [skip; type_tac; type_tac] 46 | ) 47 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/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 [inst_lift] 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/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/Theory/Typing/Stratified.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 | def DefInv (env : VEnv) (U : Nat) (Γ : List VExpr) : VExpr → VExpr → Prop 10 | | .forallE A B, .forallE A' B' => 11 | ∃ u v, env.IsDefEq U Γ A A' (.sort u) ∧ env.IsDefEq U (A::Γ) B B' (.sort v) 12 | | .forallE .., .sort .. | .sort .., .forallE .. => False 13 | | .sort u, .sort v => u ≈ v 14 | | _, _ => True 15 | 16 | variable! (henv : Ordered env) in 17 | nonrec theorem DefInv.symm (h : DefInv env U Γ e1 e2) : DefInv env U Γ e2 e1 := by 18 | cases e1 <;> cases e2 <;> try trivial 19 | · exact h.symm 20 | · let ⟨u, v, h1, h2⟩ := h; exact ⟨u, v, h1.symm, h1.defeqDF_l henv h2.symm⟩ 21 | 22 | section 23 | set_option hygiene false 24 | local notation:65 Γ " ⊢ " e " : " A:30 => HasType1 Γ e A 25 | local notation:65 Γ " ⊢ " e1 " ≡ " e2 " : " A:30 => IsDefEq1 Γ e1 e2 A 26 | 27 | variable (env : VEnv) (uvars : Nat) 28 | 29 | variable (IsDefEq1 : List VExpr → VExpr → VExpr → VExpr → Prop) in 30 | inductive HasType1 : List VExpr → VExpr → VExpr → Prop where 31 | | bvar : Lookup Γ i A → Γ ⊢ .bvar i : A 32 | | const : 33 | env.constants c = some (some ci) → 34 | (∀ l ∈ ls, l.WF uvars) → 35 | ls.length = ci.uvars → 36 | Γ ⊢ .const c ls : ci.type.instL ls 37 | | sort : l.WF uvars → Γ ⊢ .sort l : .sort (.succ l) 38 | | app : Γ ⊢ f : .forallE A B → Γ ⊢ a : A → Γ ⊢ .app f a : B.inst a 39 | | lam : Γ ⊢ A : .sort u → A::Γ ⊢ body : B → Γ ⊢ .lam A body : .forallE A B 40 | | forallE : Γ ⊢ A : .sort u → A::Γ ⊢ body : .sort v → Γ ⊢ .forallE A body : .sort (.imax u v) 41 | | defeq : Γ ⊢ A ≡ B : .sort u → Γ ⊢ e : A → Γ ⊢ e : B 42 | 43 | variable 44 | (HasType1 : List VExpr → VExpr → VExpr → Prop) 45 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) in 46 | inductive IsDefEq1 : List VExpr → VExpr → VExpr → VExpr → Prop where 47 | | refl : Γ ⊢ e : A → Γ ⊢ e ≡ e : A 48 | | symm : Γ ⊢ e ≡ e' : A → Γ ⊢ e' ≡ e : A 49 | | trans : Γ ⊢ e₁ ≡ e₂ : A → Γ ⊢ e₂ ≡ e₃ : A → Γ ⊢ e₁ ≡ e₃ : A 50 | | constDF : 51 | env.constants c = some (some ci) → 52 | (∀ l ∈ ls, l.WF uvars) → 53 | (∀ l ∈ ls', l.WF uvars) → 54 | ls.length = ci.uvars → 55 | List.Forall₂ (· ≈ ·) ls ls' → 56 | Γ ⊢ .const c ls ≡ .const c ls' : ci.type.instL ls 57 | | sortDF : l.WF uvars → l'.WF uvars → l ≈ l' → Γ ⊢ .sort l ≡ .sort l' : .sort l.succ 58 | | appDF : 59 | Γ ⊢ f ≡ f' : .forallE A B → Γ ⊢ a ≡ a' : A → Γ ⊢ .app f a ≡ .app f' a' : B.inst a 60 | | lamDF : Γ ⊢ A ≡ A' : .sort u → A::Γ ⊢ b ≡ b' : B → Γ ⊢ .lam A b ≡ .lam A' b' : .forallE A B 61 | | forallEDF : 62 | Γ ⊢ A ≡ A' : .sort u → A::Γ ⊢ B ≡ B' : .sort v → 63 | Γ ⊢ .forallE A B ≡ .forallE A' B' : .sort (.imax u v) 64 | | defeqDF : defEq Γ A B (.sort u) → Γ ⊢ e₁ ≡ e₂ : A → Γ ⊢ e₁ ≡ e₂ : B 65 | | beta : A::Γ ⊢ e : B → Γ ⊢ e' : A → Γ ⊢ .app (.lam A e) e' ≡ e.inst e' : B.inst e' 66 | | eta : Γ ⊢ e : .forallE A B → Γ ⊢ .lam A (.app e.lift (.bvar 0)) ≡ e : .forallE A B 67 | | proofIrrel : Γ ⊢ p : .sort .zero → Γ ⊢ h : p → Γ ⊢ h' : p → Γ ⊢ h ≡ h' : p 68 | | extra : 69 | env.defeqs df → (∀ l ∈ ls, l.WF uvars) → ls.length = df.uvars → 70 | Γ ⊢ df.lhs.instL ls ≡ df.rhs.instL ls : df.type.instL ls 71 | 72 | end 73 | 74 | variable! (henv : Ordered env) (hΓ : OnCtx Γ (env.IsType U)) in 75 | theorem IsDefEq.induction1 76 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) 77 | (hasType : List VExpr → VExpr → VExpr → Prop) 78 | (hty : ∀ {Γ e A}, HasType1 env U defEq Γ e A → hasType Γ e A) 79 | (hdf : ∀ {Γ e1 e2 A}, IsDefEq1 env U hasType defEq Γ e1 e2 A → defEq Γ e1 e2 A) 80 | (H : env.IsDefEq U Γ e1 e2 A) : 81 | HasType1 env U defEq Γ e1 A ∧ 82 | HasType1 env U defEq Γ e2 A ∧ 83 | IsDefEq1 env U hasType defEq Γ e1 e2 A := by 84 | have H' := H.strong henv hΓ; clear hΓ H 85 | induction H' with 86 | | bvar h => exact ⟨.bvar h, .bvar h, .refl (hty (.bvar h))⟩ 87 | | symm _ ih => exact ⟨ih.2.1, ih.1, .symm ih.2.2⟩ 88 | | trans _ _ ih1 ih2 => exact ⟨ih1.1, ih2.2.1, .trans ih1.2.2 ih2.2.2⟩ 89 | | @constDF _ _ ls₁ ls₂ u _ h1 h2 h3 h4 h5 => 90 | exact ⟨.const h1 h2 h4, 91 | .defeq (u := u.inst ls₁) sorry <| .const h1 h3 (h5.length_eq.symm.trans h4), 92 | .constDF h1 h2 h3 h4 h5⟩ 93 | | @sortDF l l' _ h1 h2 h3 => 94 | refine ⟨.sort h1, ?_, .sortDF h1 h2 h3⟩ 95 | exact .defeq (hdf <| .symm <| .sortDF (l' := l'.succ) h1 h2 (VLevel.succ_congr h3)) (.sort h2) 96 | | appDF _ _ _ _ _ _ _ _ _ ihf iha ihBa => 97 | let ⟨hf, hf', ff⟩ := ihf; let ⟨ha, ha', aa⟩ := iha 98 | exact ⟨.app hf ha, .defeq (hdf ihBa.2.2.symm) (.app hf' ha'), .appDF ff aa⟩ 99 | | lamDF _ _ _ _ _ _ _ ihA ihB _ ihb ihb' => 100 | refine ⟨.lam ihA.1 ihb.1, ?_, .lamDF ihA.2.2 ihb.2.2⟩ 101 | exact .defeq (hdf <| .symm <| .forallEDF ihA.2.2 ihB.2.2) <| .lam ihA.2.1 ihb'.2.1 102 | | forallEDF _ _ _ _ _ ih1 ih2 ih3 => 103 | exact ⟨.forallE ih1.1 ih2.1, .forallE ih1.2.1 ih3.2.1, .forallEDF ih1.2.2 ih2.2.2⟩ 104 | | defeqDF _ _ _ ih1 ih2 => 105 | exact ⟨.defeq (hdf ih1.2.2) ih2.1, .defeq (hdf ih1.2.2) ih2.2.1, .defeqDF (hdf ih1.2.2) ih2.2.2⟩ 106 | | beta _ _ _ _ _ _ _ _ ihA _ ihe ihe' _ ihee => 107 | exact ⟨.app (.lam ihA.1 ihe.1) ihe'.1, ihee.1, .beta (hty ihe.1) (hty ihe'.1)⟩ 108 | | eta _ _ _ _ _ _ _ ihA _ _ ihe ihe' => 109 | have := HasType1.app ihe'.1 (.bvar .zero) 110 | rw [instN_bvar0] at this 111 | exact ⟨.lam ihA.1 this, ihe.1, .eta (hty ihe.1)⟩ 112 | | proofIrrel _ _ _ ih1 ih2 ih3 => 113 | exact ⟨ih2.1, ih3.1, .proofIrrel (hty ih1.1) (hty ih2.1) (hty ih3.1)⟩ 114 | | extra h1 h2 h3 _ _ _ _ _ _ _ _ _ ihl' ihr' => 115 | exact ⟨ihl'.1, ihr'.1, .extra h1 h2 h3⟩ 116 | 117 | variable! {env : VEnv} 118 | {defEq : List VExpr → VExpr → VExpr → VExpr → Prop} 119 | (IH : ∀ {Γ e1 e2 A}, defEq Γ e1 e2 A → env.IsDefEq U Γ e1 e2 A) in 120 | theorem HasType1.induction (H : env.HasType1 U defEq Γ e A) : env.HasType U Γ e A := by 121 | induction H with 122 | | bvar h => exact .bvar h 123 | | const h1 h2 h3 => exact .const h1 h2 h3 124 | | sort h => exact .sort h 125 | | app _ _ ih1 ih2 => exact .app ih1 ih2 126 | | lam _ _ ih1 ih2 => exact .lam ih1 ih2 127 | | forallE _ _ ih1 ih2 => exact .forallE ih1 ih2 128 | | defeq h1 _ ih => exact (IH h1).defeq ih 129 | 130 | variable! {env : VEnv} 131 | {hasType : List VExpr → VExpr → VExpr → Prop} 132 | {defEq : List VExpr → VExpr → VExpr → VExpr → Prop} 133 | (hty : ∀ {Γ e A}, hasType Γ e A → env.HasType U Γ e A) 134 | (hdf : ∀ {Γ e1 e2 A}, defEq Γ e1 e2 A → env.IsDefEq U Γ e1 e2 A) in 135 | theorem IsDefEq1.induction 136 | (H : env.IsDefEq1 U hasType defEq Γ e1 e2 A) : env.IsDefEq U Γ e1 e2 A := by 137 | induction H with 138 | | refl h => exact hty h 139 | | symm _ ih => exact ih.symm 140 | | trans _ _ ih1 ih2 => exact ih1.trans ih2 141 | | constDF h1 h2 h3 h4 h5 => exact .constDF h1 h2 h3 h4 h5 142 | | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 143 | | appDF _ _ ih1 ih2 => exact .appDF ih1 ih2 144 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 145 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 146 | | defeqDF h1 _ ih => exact .defeqDF (hdf h1) ih 147 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 148 | | eta h => exact .eta (hty h) 149 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 150 | | extra h1 h2 h3 => exact .extra h1 h2 h3 151 | 152 | /- 153 | variable! 154 | {U : Nat} 155 | (hasType : List VExpr → VExpr → VExpr → Prop) 156 | (hasType' : List VExpr → VExpr → VExpr → Prop) 157 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) 158 | (refl : ∀ {Γ e A}, hasType Γ e A → hasType' Γ e A) 159 | (sort : ∀ {Γ l l'}, l.WF U → l'.WF U → l ≈ l' → 160 | hasType' Γ (.sort l') (.sort l.succ)) 161 | (app : ∀ {Γ f a A B}, 162 | hasType' Γ f (.forallE A B) → hasType' Γ a A → hasType' Γ (.app f a) (B.inst a)) 163 | in 164 | protected theorem IsDefEq1.hasType_ind 165 | (H : IsDefEq1 env U hasType defEq Γ e1 e2 A) : 166 | hasType' Γ e1 A ∧ hasType' Γ e2 A := by 167 | induction H with 168 | | refl h => exact ⟨refl h, refl h⟩ 169 | | symm _ ih => exact ih.symm 170 | | trans _ _ ih1 ih2 => exact ⟨ih1.1, ih2.2⟩ 171 | | sortDF h1 h2 h3 => exact ⟨sort h1 h1 rfl, sort h1 h2 h3⟩ 172 | | appDF _ _ ih1 ih2 => 173 | exact ⟨app ih1.1 ih2.1, app _ _⟩ 174 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 175 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 176 | | defeqDF h1 _ ih => exact .defeqDF (hdf h1) ih 177 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 178 | | eta h => exact .eta (hty h) 179 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 180 | | extra h1 h2 h3 => exact .extra h1 h2 h3 181 | 182 | variable! 183 | (hasType : List VExpr → VExpr → VExpr → Prop) 184 | (defEq : List VExpr → VExpr → VExpr → VExpr → Prop) in 185 | inductive IsDefEqU1 : List VExpr → VExpr → VExpr → VLevel → Prop 186 | | refl : hasType Γ A (.sort u) → IsDefEqU1 Γ A A u 187 | 188 | variable! (henv : Ordered env) 189 | {hasType : List VExpr → VExpr → VExpr → Prop} 190 | {defEq : List VExpr → VExpr → VExpr → VExpr → Prop} 191 | (refl : ∀ {Γ e A}, hasType Γ e A → defEq' Γ e e A) 192 | (hdf : ∀ {Γ e1 e2 A}, defEq Γ e1 e2 A → env.IsDefEq U Γ e1 e2 A) in 193 | theorem IsDefEq1.unique_typing1 194 | (H : env.IsDefEq1 U hasType defEq Γ e1 e2 A) : 195 | hasType Γ e1 A ∧ hasType Γ e2 A := by 196 | induction H with 197 | | refl h => exact hty h 198 | | symm _ ih => exact ih.symm 199 | | trans _ _ ih1 ih2 => exact ih1.trans ih2 200 | | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 201 | | appDF _ _ ih1 ih2 => exact .appDF ih1 ih2 202 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 203 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 204 | | defeqDF h1 _ ih => exact .defeqDF (hdf h1) ih 205 | | beta h1 h2 => exact .beta (hty h1) (hty h2) 206 | | eta h => exact .eta (hty h) 207 | | proofIrrel h1 h2 h3 => exact .proofIrrel (hty h1) (hty h2) (hty h3) 208 | | extra h1 h2 h3 => exact .extra h1 h2 h3 209 | 210 | variable! (henv : Ordered env) in 211 | theorem HasType1.unique_typing' 212 | (H1 : env.HasType1 U (IsDefEq1 env U hasType defEq) Γ e A1) 213 | (H2 : env.HasType1 U (IsDefEq1 env U hasType defEq) Γ e A2) : 214 | ∃ u, env.IsDefEq1 U hasType defEq Γ A1 A2 (.sort u) := by 215 | generalize eq1 : e = e' at H2 216 | induction H1 generalizing e' A2 with subst eq1 217 | | defeq h1 _ ih => 218 | let ⟨_, ih⟩ := ih _ rfl H2 219 | exact ⟨_, h1.trans _⟩ 220 | done 221 | | _ => ?_ 222 | <;> cases H2 223 | 224 | case bvar.bvar _ => 225 | done 226 | case const.const _ _ _ => 227 | done 228 | case sort.sort _ => 229 | done 230 | case app.app _ _ _ _ => 231 | done 232 | case lam.lam _ _ _ _ => 233 | done 234 | case forallE.forallE _ _ _ _ => 235 | done 236 | case defeq.defeq _ _ _ => 237 | done 238 | _ 239 | -- | bvar h => 240 | -- refine .bvar h.instL 241 | -- | @const _ _ ls' _ h1 h2 h3 => 242 | -- simp [instL, instL_instL] 243 | -- exact .const h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 244 | -- | sort _ _ h3 => 245 | -- exact .sortDF (VLevel.WF.inst hls) (VLevel.WF.inst hls) (VLevel.inst_congr_l h3) 246 | -- | app _ _ ih1 ih2 => exact instL_instN ▸ .appDF ih1 ih2 247 | -- | lam _ _ ih1 ih2 => exact .lamDF ih1 ih2 248 | -- | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 249 | -- | defeqDF _ _ ih1 ih2 => exact .defeqDF ih1 ih2 250 | -- | beta _ _ ih1 ih2 => simpa using .beta ih1 ih2 251 | -- | eta _ ih => simpa [instL] using .eta ih 252 | -- | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel ih1 ih2 ih3 253 | -- | extra h1 h2 h3 => 254 | -- simp [instL, instL_instL] 255 | -- exact .extra h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 256 | -/ 257 | 258 | /- 259 | variable! (henv : Ordered env) in 260 | theorem IsDefEq.unique_typing' 261 | (H1 : env.IsDefEq U Γ e1 e2 A1) (H2 : env.IsDefEq U Γ e1 e2 A2) : 262 | ∃ u, env.IsDefEq U Γ A1 A2 (.sort u) := by 263 | generalize eq1 : e1 = e1', eq2 : e2 = e2' at H2 264 | induction H1 generalizing A2 with 265 | | bvar h => 266 | 267 | refine .bvar h.instL 268 | | @const _ _ ls' _ h1 h2 h3 => 269 | simp [instL, instL_instL] 270 | exact .const h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 271 | | symm _ ih => exact .symm ih 272 | | trans _ _ ih1 ih2 => exact .trans ih1 ih2 273 | | sortDF _ _ h3 => 274 | exact .sortDF (VLevel.WF.inst hls) (VLevel.WF.inst hls) (VLevel.inst_congr_l h3) 275 | | appDF _ _ ih1 ih2 => exact instL_instN ▸ .appDF ih1 ih2 276 | | lamDF _ _ ih1 ih2 => exact .lamDF ih1 ih2 277 | | forallEDF _ _ ih1 ih2 => exact .forallEDF ih1 ih2 278 | | defeqDF _ _ ih1 ih2 => exact .defeqDF ih1 ih2 279 | | beta _ _ ih1 ih2 => simpa using .beta ih1 ih2 280 | | eta _ ih => simpa [instL] using .eta ih 281 | | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel ih1 ih2 ih3 282 | | extra h1 h2 h3 => 283 | simp [instL, instL_instL] 284 | exact .extra h1 (by simp [h2, VLevel.WF.inst hls]) (by simp [h3]) 285 | -/ 286 | 287 | 288 | /- depends on church-rosser 289 | variable! {env : VEnv} (henv : env.Ordered) in 290 | theorem IsDefEq.weakN_inv (W : Ctx.LiftN n k Γ Γ') 291 | (H : env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k)) : 292 | env.IsDefEq U Γ e1 e2 A := by 293 | generalize eq1 : e1.liftN n k = e1', eq2 : e2.liftN n k = e2', eqA : A.liftN n k = A' at H 294 | induction H generalizing k e1 e2 A with 295 | | bvar h => 296 | cases eqA; cases e1 <;> cases eq1; cases e2 <;> injection eq2 297 | cases liftVar_inj.1 ‹_›; exact .bvar (h.weakN_inv W) 298 | | @const c ci ls Γ h1 h2 h3 => 299 | cases e1 <;> cases eq1; cases e2 <;> cases eq2 300 | rw [ClosedN.liftN_eq_rev (eqA ▸ (henv.closedC h1).instL) (Nat.zero_le _)] at eqA 301 | exact eqA ▸ .const h1 h2 h3 302 | | symm _ ih => exact .symm (ih W eq2 eq1 eqA) 303 | | trans _ _ ih1 ih2 => sorry 304 | -- | sortDF h1 h2 h3 => exact .sortDF h1 h2 h3 305 | -- | appDF _ _ ih1 ih2 => exact liftN_inst_hi .. ▸ .appDF (ih1 W) (ih2 W) 306 | -- | lamDF _ _ ih1 ih2 => exact .lamDF (ih1 W) (ih2 W.succ) 307 | -- | forallEDF _ _ ih1 ih2 => exact .forallEDF (ih1 W) (ih2 W.succ) 308 | -- | defeqDF _ _ ih1 ih2 => exact .defeqDF (ih1 W) (ih2 W) 309 | -- | beta _ _ ih1 ih2 => 310 | -- exact liftN_inst_hi .. ▸ liftN_instN_hi .. ▸ .beta (ih1 W.succ) (ih2 W) 311 | -- | eta _ ih => 312 | -- have := IsDefEq.eta (ih W) 313 | -- simp [liftN]; rwa [← lift_liftN'] 314 | -- | proofIrrel _ _ _ ih1 ih2 ih3 => exact .proofIrrel (ih1 W) (ih2 W) (ih3 W) 315 | -- | extra h1 h2 h3 => 316 | -- have ⟨⟨hA1, _⟩, hA2, hA3⟩ := henv.closed.2 h1 317 | -- rw [ 318 | -- hA1.instL.liftN_eq (Nat.zero_le _), 319 | -- hA2.instL.liftN_eq (Nat.zero_le _), 320 | -- hA3.instL.liftN_eq (Nat.zero_le _)] 321 | -- exact .extra h1 h2 h3 322 | | _ => sorry 323 | 324 | variable! {env : VEnv} (henv : env.Ordered) in 325 | theorem HasType.weakN_inv (W : Ctx.LiftN n k Γ Γ') 326 | (H : env.HasType U Γ' (e.liftN n k) (A.liftN n k)) : 327 | env.HasType U Γ e A := IsDefEq.weakN_inv henv W H 328 | 329 | variable! {env : VEnv} (henv : env.Ordered) in 330 | theorem IsType.weakN_inv (W : Ctx.LiftN n k Γ Γ') (H : env.IsType U Γ' (A.liftN n k)) : 331 | env.IsType U Γ A := let ⟨_, h⟩ := H; ⟨_, h.weakN_inv henv W⟩ 332 | -/ 333 | -------------------------------------------------------------------------------- /Lean4Lean/Theory/Typing/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/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 | theorem isDefEq_iff (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) : 24 | env.IsDefEq U Γ e₁ e₂ A ↔ 25 | env.HasType U Γ e₁ A ∧ env.HasType U Γ e₂ A ∧ env.IsDefEqU U Γ e₁ e₂ := by 26 | refine ⟨fun h => ⟨h.hasType.1, h.hasType.2, _, h⟩, fun ⟨_, h2, _, h3⟩ => ?_⟩ 27 | have ⟨_, h⟩ := h3.uniq henv hΓ h2 28 | exact h.defeqDF h3 29 | 30 | theorem IsDefEq.trans_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 31 | (h₁ : env.IsDefEq U Γ e₁ e₂ A) (h₂ : env.IsDefEq U Γ e₂ e₃ B) : 32 | env.IsDefEq U Γ e₁ e₃ B := by 33 | have ⟨_, h⟩ := h₁.uniq henv hΓ h₂ 34 | exact .trans (.defeqDF h h₁) h₂ 35 | 36 | theorem IsDefEq.trans_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 37 | (h₁ : env.IsDefEq U Γ e₁ e₂ A) (h₂ : env.IsDefEq U Γ e₂ e₃ B) : 38 | env.IsDefEq U Γ e₁ e₃ A := by 39 | have ⟨_, h⟩ := h₁.uniq henv hΓ h₂ 40 | exact h₁.trans (.defeqDF (.symm h) h₂) 41 | 42 | theorem IsDefEqU.defeqDF (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 43 | (h₁ : env.IsDefEqU U Γ A B) (h₂ : env.IsDefEq U Γ e₁ e₂ A) : 44 | env.IsDefEq U Γ e₁ e₂ B := by 45 | have ⟨_, h₁⟩ := h₁ 46 | have ⟨_, hA⟩ := h₂.isType henv hΓ 47 | exact .defeqDF (hA.trans_l henv hΓ h₁) h₂ 48 | 49 | theorem IsDefEqU.of_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 50 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.HasType U Γ e₁ A) : 51 | env.IsDefEq U Γ e₁ e₂ A := let ⟨_, h⟩ := h1; h2.trans_l henv hΓ h 52 | 53 | theorem HasType.defeqU_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.HasType U Γ e₂ A := (h1.of_l henv hΓ h2).hasType.2 56 | 57 | theorem IsType.defeqU_l (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 58 | (h1 : env.IsDefEqU U Γ A₁ A₂) (h2 : env.IsType U Γ A₁) : 59 | env.IsType U Γ A₂ := h2.imp fun _ h2 => h2.defeqU_l henv hΓ h1 60 | 61 | theorem IsDefEqU.of_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 62 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.HasType U Γ e₂ A) : 63 | env.IsDefEq U Γ e₁ e₂ A := (h1.symm.of_l henv hΓ h2).symm 64 | 65 | theorem HasType.defeqU_r (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 66 | (h1 : env.IsDefEqU U Γ A₁ A₂) (h2 : env.HasType U Γ e A₁) : 67 | env.HasType U Γ e A₂ := h1.defeqDF henv hΓ h2 68 | 69 | theorem IsDefEqU.trans (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 70 | (h1 : env.IsDefEqU U Γ e₁ e₂) (h2 : env.IsDefEqU U Γ e₂ e₃) : 71 | env.IsDefEqU U Γ e₁ e₃ := h1.imp fun _ h1 => let ⟨_, h2⟩ := h2; h1.trans_l henv hΓ h2 72 | 73 | theorem IsDefEqU.sort_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 74 | (h1 : env.IsDefEqU U Γ (.sort u) (.sort v)) : u ≈ v := sorry 75 | 76 | theorem IsDefEqU.forallE_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) 77 | (h1 : env.IsDefEqU U Γ (.forallE A B) (.forallE A' B')) : 78 | (∃ u, env.IsDefEq U Γ A A' (.sort u)) ∧ ∃ u, env.IsDefEq U (A::Γ) B B' (.sort u) := sorry 79 | 80 | theorem IsDefEqU.sort_forallE_inv (henv : VEnv.WF env) (hΓ : OnCtx Γ (env.IsType U)) : 81 | ¬env.IsDefEqU U Γ (.sort u) (.forallE A B) := sorry 82 | 83 | variable! (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) in 84 | theorem IsDefEqU.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 85 | env.IsDefEqU U Γ' (e1.liftN n k) (e2.liftN n k) ↔ env.IsDefEqU U Γ e1 e2 := by 86 | refine ⟨fun h => have := henv; have := hΓ; sorry, fun h => h.weakN henv W⟩ 87 | 88 | variable! (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) in 89 | theorem _root_.Lean4Lean.VExpr.WF.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 90 | VExpr.WF env U Γ' (e.liftN n k) ↔ VExpr.WF env U Γ e := IsDefEqU.weakN_iff henv hΓ W 91 | 92 | theorem IsDefEq.skips (henv : VEnv.WF env) (hΓ : OnCtx Γ' (env.IsType U)) 93 | (W : Ctx.LiftN n k Γ Γ') 94 | (H : env.IsDefEq U Γ' e₁ e₂ A) (h1 : e₁.Skips n k) (h2 : e₂.Skips n k) : 95 | ∃ B, env.IsDefEq U Γ' e₁ e₂ B ∧ B.Skips n k := by 96 | obtain ⟨e₁, rfl⟩ := VExpr.skips_iff_exists.1 h1 97 | obtain ⟨e₂, rfl⟩ := VExpr.skips_iff_exists.1 h2 98 | have ⟨_, H⟩ := (IsDefEqU.weakN_iff henv hΓ W).1 ⟨_, H⟩ 99 | exact ⟨_, H.weakN henv W, .liftN⟩ 100 | 101 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) (hΓ : OnCtx Γ (env.IsType U)) in 102 | theorem IsDefEq.weakN_iff' (W : Ctx.LiftN n k Γ Γ') : 103 | env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k) ↔ env.IsDefEq U Γ e1 e2 A := by 104 | refine ⟨fun h => ?_, fun h => h.weakN henv W⟩ 105 | have ⟨_, H⟩ := (IsDefEqU.weakN_iff henv hΓ' W).1 ⟨_, h⟩ 106 | refine IsDefEqU.defeqDF henv hΓ ?_ H 107 | exact (IsDefEqU.weakN_iff henv hΓ' W).1 <| (H.weakN henv W).uniqU henv hΓ' h.symm 108 | 109 | variable! (henv : VEnv.WF env) in 110 | theorem _root_.Lean4Lean.OnCtx.weakN_inv 111 | (W : Ctx.LiftN n k Γ Γ') (H : OnCtx Γ' (env.IsType U)) : OnCtx Γ (env.IsType U) := by 112 | induction W with 113 | | zero As h => 114 | clear h 115 | induction As with 116 | | nil => exact H 117 | | cons A As ih => exact ih H.1 118 | | succ W ih => 119 | let ⟨H1, _, H2⟩ := H 120 | exact ⟨ih H1, _, (IsDefEq.weakN_iff' henv H1 (ih H1) W).1 H2⟩ 121 | 122 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 123 | theorem IsDefEq.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 124 | env.IsDefEq U Γ' (e1.liftN n k) (e2.liftN n k) (A.liftN n k) ↔ env.IsDefEq U Γ e1 e2 A := 125 | IsDefEq.weakN_iff' henv hΓ' (hΓ'.weakN_inv henv W) W 126 | 127 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 128 | theorem HasType.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 129 | env.HasType U Γ' (e.liftN n k) (A.liftN n k) ↔ env.HasType U Γ e A := 130 | IsDefEq.weakN_iff henv hΓ' W 131 | 132 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 133 | theorem IsType.weakN_iff (W : Ctx.LiftN n k Γ Γ') : 134 | env.IsType U Γ' (A.liftN n k) ↔ env.IsType U Γ A := 135 | exists_congr fun _ => HasType.weakN_iff henv hΓ' W (A := .sort _) 136 | 137 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 138 | theorem HasType.skips (W : Ctx.LiftN n k Γ Γ') 139 | (h1 : env.HasType U Γ' e A) (h2 : e.Skips n k) : ∃ B, env.HasType U Γ' e B ∧ B.Skips n k := 140 | IsDefEq.skips henv hΓ' W h1 h2 h2 141 | 142 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 143 | theorem IsDefEqU.weak'_iff (W : Ctx.Lift' l Γ Γ') : 144 | env.IsDefEqU U Γ' (e1.lift' l) (e2.lift' l) ↔ env.IsDefEqU U Γ e1 e2 := by 145 | generalize e : l.depth = n 146 | induction n generalizing l Γ' with 147 | | zero => simp [VExpr.lift'_depth_zero e, W.depth_zero e] 148 | | succ n ih => 149 | obtain ⟨l, k, rfl, rfl⟩ := Lift.depth_succ e 150 | have ⟨Γ₁, W1, W2⟩ := W.of_cons_skip 151 | simp only [Lift.consN_skip_eq, VExpr.lift'_comp, VExpr.lift'_consN_skipN, 152 | weakN_iff henv hΓ' W2, ih (hΓ'.weakN_inv henv W2) W1 Lift.depth_consN] 153 | 154 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 155 | theorem IsDefEq.weak'_iff (W : Ctx.Lift' l Γ Γ') : 156 | env.IsDefEq U Γ' (e1.lift' l) (e2.lift' l) (A.lift' l) ↔ env.IsDefEq U Γ e1 e2 A := by 157 | generalize e : l.depth = n 158 | induction n generalizing l Γ' with 159 | | zero => simp [VExpr.lift'_depth_zero e, W.depth_zero e] 160 | | succ n ih => 161 | obtain ⟨l, k, rfl, rfl⟩ := Lift.depth_succ e 162 | have ⟨Γ₁, W1, W2⟩ := W.of_cons_skip 163 | simp only [Lift.consN_skip_eq, VExpr.lift'_comp, VExpr.lift'_consN_skipN, 164 | weakN_iff henv hΓ' W2, ih (hΓ'.weakN_inv henv W2) W1 Lift.depth_consN] 165 | 166 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 167 | theorem HasType.weak'_iff (W : Ctx.Lift' l Γ Γ') : 168 | env.HasType U Γ' (e.lift' l) (A.lift' l) ↔ env.HasType U Γ e A := 169 | IsDefEq.weak'_iff henv hΓ' W 170 | 171 | variable! (henv : VEnv.WF env) (hΓ' : OnCtx Γ' (env.IsType U)) in 172 | theorem IsType.weak'_iff (W : Ctx.Lift' l Γ Γ') : 173 | env.IsType U Γ' (e.lift' l) ↔ env.IsType U Γ e := 174 | exists_congr fun _ => HasType.weak'_iff henv hΓ' W (A := .sort _) 175 | -------------------------------------------------------------------------------- /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 ((List.range v.uvars).map .param), 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/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 | 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.addConst (env : VEnv) (name : Name) (ci : Option VConstant) : Option VEnv := 26 | match env.constants name with 27 | | some _ => none 28 | | none => some { env with constants := fun n => if name = n then some ci else env.constants n } 29 | 30 | def VEnv.addDefEq (env : VEnv) (df : VDefEq) : VEnv := 31 | { env with defeqs := fun x => x = df ∨ env.defeqs x } 32 | 33 | structure VEnv.LE (env1 env2 : VEnv) : Prop where 34 | constants : env1.constants n = some a → env2.constants n = some a 35 | defeqs : env1.defeqs df → env2.defeqs df 36 | 37 | instance : LE VEnv := ⟨VEnv.LE⟩ 38 | 39 | theorem VEnv.LE.rfl {env : VEnv} : env ≤ env := ⟨id, id⟩ 40 | 41 | theorem VEnv.LE.trans {a b c : VEnv} (h1 : a ≤ b) (h2 : b ≤ c) : a ≤ c := 42 | ⟨h2.1 ∘ h1.1, h2.2 ∘ h1.2⟩ 43 | -------------------------------------------------------------------------------- /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 le_antisymm_iff {a b : VLevel} : a ≈ b ↔ a ≤ b ∧ b ≤ a := 63 | equiv_def.trans <| (forall_congr' fun _ => Nat.le_antisymm_iff).trans forall_and 64 | 65 | theorem succ_congr {a b : VLevel} (h : a ≈ b) : succ a ≈ succ b := by 66 | simpa [equiv_def, eval] using h 67 | 68 | theorem max_congr (h₁ : a₁ ≈ b₁) (h₂ : a₂ ≈ b₂) : max a₁ a₂ ≈ max b₁ b₂ := by 69 | simp_all [equiv_def, eval] 70 | 71 | theorem imax_congr (h₁ : a₁ ≈ b₁) (h₂ : a₂ ≈ b₂) : imax a₁ a₂ ≈ imax b₁ b₂ := by 72 | simp_all [equiv_def, eval] 73 | 74 | theorem max_comm : max a b ≈ max b a := by simp [equiv_def, eval, Nat.max_comm] 75 | 76 | theorem LE.max_eq_left (h : b.LE a) : max a b ≈ a := by 77 | simp [equiv_def, eval, Nat.max_eq_left (h _)] 78 | 79 | theorem LE.max_eq_right (h : a.LE b) : max a b ≈ b := by 80 | simp [equiv_def, eval, Nat.max_eq_right (h _)] 81 | 82 | theorem max_self : max a a ≈ a := by simp [equiv_def, eval] 83 | 84 | theorem zero_imax : imax zero a ≈ a := by 85 | simp [equiv_def, eval, Nat.imax, eq_comm (b := 0)] 86 | 87 | theorem imax_self : imax a a ≈ a := by 88 | simp [equiv_def, eval, Nat.imax, eq_comm (b := 0)] 89 | 90 | def IsNeverZero (a : VLevel) : Prop := ∀ ls, a.eval ls ≠ 0 91 | 92 | theorem IsNeverZero.imax_eq_max (h : IsNeverZero b) : imax a b ≈ max a b := by 93 | simp_all [equiv_def, eval, Nat.imax, IsNeverZero] 94 | 95 | variable (ls : List VLevel) in 96 | def inst : VLevel → VLevel 97 | | .zero => .zero 98 | | .succ l => .succ l.inst 99 | | .max l₁ l₂ => .max l₁.inst l₂.inst 100 | | .imax l₁ l₂ => .imax l₁.inst l₂.inst 101 | | .param i => ls.getD i .zero 102 | 103 | theorem inst_inst {l : VLevel} : (l.inst ls).inst ls' = l.inst (ls.map (inst ls')) := by 104 | induction l <;> simp [inst, *, List.getD_eq_getElem?_getD, List.getElem?_map] 105 | case param n => cases ls[n]? <;> simp [inst] 106 | 107 | theorem inst_id {l : VLevel} (h : l.WF u) : l.inst ((List.range u).map .param) = l := by 108 | induction l <;> simp_all [inst, WF, List.getD_eq_getElem?_getD, List.getElem?_range] 109 | 110 | theorem eval_inst {l : VLevel} : (l.inst ls).eval ns = l.eval (ls.map (eval ns)) := by 111 | induction l <;> simp [eval, inst, *, List.getD_eq_getElem?_getD] 112 | case param n => cases ls[n]? <;> simp [eval] 113 | 114 | theorem WF.inst {l : VLevel} (H : ∀ l ∈ ls, l.WF n) : (l.inst ls).WF n := by 115 | induction l with 116 | | zero => trivial 117 | | succ _ ih => exact ih 118 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => exact ⟨ih1, ih2⟩ 119 | | param i => 120 | simp [VLevel.inst, List.getD_eq_getElem?_getD] 121 | cases e : ls[i]? with 122 | | none => trivial 123 | | some => exact H _ (List.mem_of_getElem? e) 124 | 125 | theorem id_WF : ∀ l ∈ (List.range u).map param, l.WF u := by simp [WF] 126 | 127 | theorem inst_congr {l : VLevel} (h1 : l ≈ l') (h2 : List.Forall₂ (·≈·) ls ls') : 128 | l.inst ls ≈ l'.inst ls' := by 129 | simp [equiv_def, eval_inst, ← equiv_def.1 h1] 130 | intro ns; congr 1 131 | induction h2 with 132 | | nil => rfl 133 | | cons h2 => simp [*, equiv_def.1 h2] 134 | 135 | theorem inst_congr_l {l : VLevel} (h1 : l ≈ l') : l.inst ls ≈ l'.inst ls := 136 | inst_congr h1 <| .rfl fun _ _ => rfl 137 | 138 | variable (ls : List Name) in 139 | def ofLevel : Lean.Level → Option VLevel 140 | | .zero => return .zero 141 | | .succ l => return .succ (← ofLevel l) 142 | | .max l₁ l₂ => return .max (← ofLevel l₁) (← ofLevel l₂) 143 | | .imax l₁ l₂ => return .imax (← ofLevel l₁) (← ofLevel l₂) 144 | | .param n => 145 | let i := ls.idxOf n 146 | if i < ls.length then some (.param i) else none 147 | | .mvar _ => none 148 | 149 | theorem WF.of_ofLevel (h : ofLevel ls l = some l') : l'.WF ls.length := by 150 | induction l generalizing l' with simp [ofLevel, bind, Option.bind_eq_some_iff] at h 151 | | zero => cases h; trivial 152 | | succ _ ih => obtain ⟨l', h, ⟨⟩⟩ := h; exact @ih l' h 153 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => obtain ⟨_, h1, _, h2, ⟨⟩⟩ := h; exact ⟨ih1 h1, ih2 h2⟩ 154 | | param n => exact h.2 ▸ h.1 155 | 156 | theorem WF.of_mapM_ofLevel (h : List.mapM (VLevel.ofLevel Us) us = some us') 157 | (a) (hl : a ∈ us') : VLevel.WF Us.length a := by 158 | rw [List.mapM_eq_some] at h 159 | have ⟨_, _, h⟩ := h.forall_exists_r _ hl; exact .of_ofLevel h 160 | -------------------------------------------------------------------------------- /Lean4Lean/Verify.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Verify.Typing.Lemmas 2 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Axioms.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Tactic.OpenPrivate 2 | import Lean4Lean.Std.Basic 3 | import Lean4Lean.Std.NodupKeys 4 | 5 | open scoped List 6 | namespace Lean 7 | 8 | noncomputable def PersistentArrayNode.toList' : PersistentArrayNode α → List α := 9 | PersistentArrayNode.rec 10 | (motive_1 := fun _ => List α) (motive_2 := fun _ => List α) (motive_3 := fun _ => List α) 11 | (node := fun _ => id) (leaf := (·.toList)) (fun _ => id) [] (fun _ _ a b => a ++ b) 12 | 13 | namespace PersistentArray 14 | 15 | inductive WF : PersistentArray α → Prop where 16 | | empty : WF .empty 17 | | push : WF arr → WF (arr.push x) 18 | 19 | noncomputable def toList' (arr : PersistentArray α) : List α := 20 | arr.root.toList' ++ arr.tail.toList 21 | 22 | @[simp] theorem toList'_empty : (.empty : PersistentArray α).toList' = [] := rfl 23 | 24 | /-- We cannot prove this because `insertNewLeaf` is partial -/ 25 | @[simp] axiom toList'_push {α} (arr : PersistentArray α) (x : α) : 26 | (arr.push x).toList' = arr.toList' ++ [x] 27 | 28 | @[simp] theorem size_empty : (.empty : PersistentArray α).size = 0 := rfl 29 | 30 | @[simp] theorem size_push {α} (arr : PersistentArray α) (x : α) : 31 | (arr.push x).size = arr.size + 1 := by 32 | simp [push]; split <;> [rfl; (simp [mkNewTail]; split <;> rfl)] 33 | 34 | @[simp] theorem WF.toList'_length (h : WF arr) : arr.toList'.length = arr.size := by 35 | induction h <;> simp [*] 36 | 37 | end PersistentArray 38 | 39 | namespace PersistentHashMap 40 | 41 | noncomputable def Node.toList' : Node α β → List (α × β) := 42 | Node.rec 43 | (motive_1 := fun _ => List (α × β)) (motive_2 := fun _ => List (α × β)) 44 | (motive_3 := fun _ => List (α × β)) (motive_4 := fun _ => List (α × β)) 45 | (entries := fun _ => id) (collision := fun ks xs _ => ks.toList.zip xs.toList) 46 | (mk := fun _ => id) 47 | (nil := []) (cons := fun _ _ l1 l2 => l1 ++ l2) 48 | (entry := fun a b => [(a, b)]) (ref := fun _ => id) (null := []) 49 | 50 | noncomputable def toList' [BEq α] [Hashable α] (m : PersistentHashMap α β) : 51 | List (α × β) := m.root.toList' 52 | 53 | inductive WF [BEq α] [Hashable α] : PersistentHashMap α β → Prop where 54 | | empty : WF .empty 55 | | insert : WF m → WF (m.insert a b) 56 | 57 | /-- We can't prove this because `Lean.PersistentHashMap.insertAux` is opaque -/ 58 | axiom WF.toList'_insert {α β} [BEq α] [Hashable α] 59 | [PartialEquivBEq α] [LawfulHashable α] 60 | {m : PersistentHashMap α β} (_ : WF m) (a : α) (b : β) : 61 | (m.insert a b).toList' ~ (a, b) :: m.toList'.filter (¬a == ·.1) 62 | 63 | /-- We can't prove this because `Lean.PersistentHashMap.findAux` is opaque -/ 64 | axiom WF.find?_eq {α β} [BEq α] [Hashable α] 65 | [PartialEquivBEq α] [LawfulHashable α] 66 | {m : PersistentHashMap α β} (_ : WF m) (a : α) : m.find? a = m.toList'.lookup a 67 | 68 | /-- We can't prove this because `Lean.PersistentHashMap.{findAux, containsAux}` are opaque -/ 69 | axiom findAux_isSome {α β} [BEq α] {node : Node α β} (i : USize) (a : α) : 70 | containsAux node i a = (findAux node i a).isSome 71 | 72 | end PersistentHashMap 73 | 74 | -- FIXME: lean4#8464 75 | open private mkAppRangeAux from Lean.Expr in 76 | axiom Expr.mkAppRangeAux.eq_def (n : Nat) (args : Array Expr) (i : Nat) (e : Expr) : 77 | mkAppRangeAux n args i e = 78 | if i < n then mkAppRangeAux n args (i + 1) (mkApp e args[i]!) else e 79 | 80 | namespace Expr 81 | 82 | def hasFVar' : Expr → Bool 83 | | .fvar _ => true 84 | | .const .. 85 | | .bvar _ 86 | | .sort _ 87 | | .mvar _ 88 | | .lit _ => false 89 | | .mdata _ e => e.hasFVar' 90 | | .proj _ _ e => e.hasFVar' 91 | | .app e1 e2 92 | | .lam _ e1 e2 _ 93 | | .forallE _ e1 e2 _ => e1.hasFVar' || e2.hasFVar' 94 | | .letE _ t v b _ => t.hasFVar' || v.hasFVar' || b.hasFVar' 95 | 96 | /-- This is currently false, see bug lean4#8554 -/ 97 | axiom hasFVar_eq (e : Expr) : e.hasFVar = e.hasFVar' 98 | 99 | def hasExprMVar' : Expr → Bool 100 | | .mvar _ => true 101 | | .const .. 102 | | .bvar _ 103 | | .sort _ 104 | | .fvar _ 105 | | .lit _ => false 106 | | .mdata _ e => e.hasExprMVar' 107 | | .proj _ _ e => e.hasExprMVar' 108 | | .app e1 e2 109 | | .lam _ e1 e2 _ 110 | | .forallE _ e1 e2 _ => e1.hasExprMVar' || e2.hasExprMVar' 111 | | .letE _ t v b _ => t.hasExprMVar' || v.hasExprMVar' || b.hasExprMVar' 112 | 113 | /-- This is currently false, see bug lean4#8554 -/ 114 | @[simp] axiom hasExprMVar_eq (e : Expr) : e.hasExprMVar = e.hasExprMVar' 115 | 116 | def hasLevelMVar' : Expr → Bool 117 | | .const _ ls => ls.any (·.hasMVar) 118 | | .sort u => u.hasMVar 119 | | .bvar _ 120 | | .fvar _ 121 | | .mvar _ 122 | | .lit _ => false 123 | | .mdata _ e => e.hasLevelMVar' 124 | | .proj _ _ e => e.hasLevelMVar' 125 | | .app e1 e2 126 | | .lam _ e1 e2 _ 127 | | .forallE _ e1 e2 _ => e1.hasLevelMVar' || e2.hasLevelMVar' 128 | | .letE _ t v b _ => t.hasLevelMVar' || v.hasLevelMVar' || b.hasLevelMVar' 129 | 130 | /-- This is currently false, see bug lean4#8554 -/ 131 | @[simp] axiom hasLevelMVar_eq (e : Expr) : e.hasLevelMVar = e.hasLevelMVar' 132 | 133 | def hasLevelParam' : Expr → Bool 134 | | .const _ ls => ls.any (·.hasParam) 135 | | .sort u => u.hasParam 136 | | .bvar _ 137 | | .fvar _ 138 | | .mvar _ 139 | | .lit _ => false 140 | | .mdata _ e => e.hasLevelParam' 141 | | .proj _ _ e => e.hasLevelParam' 142 | | .app e1 e2 143 | | .lam _ e1 e2 _ 144 | | .forallE _ e1 e2 _ => e1.hasLevelParam' || e2.hasLevelParam' 145 | | .letE _ t v b _ => t.hasLevelParam' || v.hasLevelParam' || b.hasLevelParam' 146 | 147 | /-- This is currently false, see bug lean4#8554 -/ 148 | @[simp] axiom hasLevelParam_eq (e : Expr) : e.hasLevelParam = e.hasLevelParam' 149 | 150 | def looseBVarRange' : Expr → Nat 151 | | .bvar i => i + 1 152 | | .const .. 153 | | .sort _ 154 | | .fvar _ 155 | | .mvar _ 156 | | .lit _ => 0 157 | | .mdata _ e 158 | | .proj _ _ e => e.looseBVarRange' 159 | | .app e1 e2 => max e1.looseBVarRange' e2.looseBVarRange' 160 | | .lam _ e1 e2 _ 161 | | .forallE _ e1 e2 _ => max e1.looseBVarRange' (e2.looseBVarRange' - 1) 162 | | .letE _ e1 e2 e3 _ => max (max e1.looseBVarRange' e2.looseBVarRange') (e3.looseBVarRange' - 1) 163 | 164 | /-- This is currently false, see bug lean4#8554 -/ 165 | @[simp] axiom looseBVarRange_eq (e : Expr) : e.looseBVarRange = e.looseBVarRange' 166 | 167 | /-- This could be an `@[implemented_by]` -/ 168 | @[simp] axiom replace_eq (e : Expr) (f) : e.replace f = e.replaceNoCache f 169 | 170 | def liftLooseBVars' (e : @& Expr) (s d : @& Nat) : Expr := 171 | match e with 172 | | .bvar i => .bvar (if i < s then i else i + d) 173 | | .mdata m e => .mdata m (liftLooseBVars' e s d) 174 | | .proj n i e => .proj n i (liftLooseBVars' e s d) 175 | | .app f a => .app (liftLooseBVars' f s d) (liftLooseBVars' a s d) 176 | | .lam n t b bi => .lam n (liftLooseBVars' t s d) (liftLooseBVars' b (s+1) d) bi 177 | | .forallE n t b bi => .forallE n (liftLooseBVars' t s d) (liftLooseBVars' b (s+1) d) bi 178 | | .letE n t v b bi => 179 | .letE n (liftLooseBVars' t s d) (liftLooseBVars' v s d) (liftLooseBVars' b (s+1) d) bi 180 | | e@(.const ..) 181 | | e@(.sort _) 182 | | e@(.fvar _) 183 | | e@(.mvar _) 184 | | e@(.lit _) => e 185 | 186 | /-- This could be an `@[implemented_by]` -/ 187 | @[simp] axiom liftLooseBVars_eq (e : Expr) (s d) : e.liftLooseBVars s d = e.liftLooseBVars' s d 188 | 189 | def lowerLooseBVars' (e : @& Expr) (s d : @& Nat) : Expr := 190 | if s < d then e else 191 | match e with 192 | | .bvar i => .bvar (if i < s then i else i - d) 193 | | .mdata m e => .mdata m (lowerLooseBVars' e s d) 194 | | .proj n i e => .proj n i (lowerLooseBVars' e s d) 195 | | .app f a => .app (lowerLooseBVars' f s d) (lowerLooseBVars' a s d) 196 | | .lam n t b bi => .lam n (lowerLooseBVars' t s d) (lowerLooseBVars' b (s+1) d) bi 197 | | .forallE n t b bi => .forallE n (lowerLooseBVars' t s d) (lowerLooseBVars' b (s+1) d) bi 198 | | .letE n t v b bi => 199 | .letE n (lowerLooseBVars' t s d) (lowerLooseBVars' v s d) (lowerLooseBVars' b (s+1) d) bi 200 | | e@(.const ..) 201 | | e@(.sort _) 202 | | e@(.fvar _) 203 | | e@(.mvar _) 204 | | e@(.lit _) => e 205 | 206 | /-- This could be an `@[implemented_by]` -/ 207 | @[simp] axiom lowerLooseBVars_eq (e : Expr) (s d) : e.lowerLooseBVars s d = e.lowerLooseBVars' s d 208 | 209 | def instantiate1' (e : Expr) (subst : Expr) (d := 0) : Expr := 210 | match e with 211 | | .bvar i => if i < d then e else if i = d then subst.liftLooseBVars' 0 d else .bvar (i - 1) 212 | | .mdata m e => .mdata m (instantiate1' e subst d) 213 | | .proj s i e => .proj s i (instantiate1' e subst d) 214 | | .app f a => .app (instantiate1' f subst d) (instantiate1' a subst d) 215 | | .lam n t b bi => .lam n (instantiate1' t subst d) (instantiate1' b subst (d+1)) bi 216 | | .forallE n t b bi => .forallE n (instantiate1' t subst d) (instantiate1' b subst (d+1)) bi 217 | | .letE n t v b bi => 218 | .letE n (instantiate1' t subst d) (instantiate1' v subst d) (instantiate1' b subst (d+1)) bi 219 | | .const .. 220 | | .sort _ 221 | | .fvar _ 222 | | .mvar _ 223 | | .lit _ => e 224 | 225 | /-- This could be an `@[implemented_by]` -/ 226 | @[simp] axiom instantiate1_eq (e : Expr) (subst) : e.instantiate1 subst = e.instantiate1' subst 227 | 228 | @[simp] def instantiateList : Expr → List Expr → (k :_:= 0) → Expr 229 | | e, [], _ => e 230 | | e, a :: as, k => instantiateList (instantiate1' e a k) as k 231 | 232 | /-- This could be an `@[implemented_by]` -/ 233 | @[simp] axiom instantiate_eq (e : Expr) (subst) : 234 | e.instantiate subst = e.instantiateList subst.toList 235 | 236 | /-- This could be an `@[implemented_by]` -/ 237 | @[simp] axiom instantiateRev_eq (e : Expr) (subst) : 238 | e.instantiateRev subst = e.instantiate subst.reverse 239 | 240 | /-- This could be an `@[implemented_by]` -/ 241 | @[simp] axiom instantiateRange_eq (e : Expr) (subst) : 242 | e.instantiateRange start stop subst = e.instantiate (subst.extract start stop) 243 | 244 | /-- This could be an `@[implemented_by]` -/ 245 | @[simp] axiom instantiateRevRange_eq (e : Expr) (subst) : 246 | e.instantiateRevRange start stop subst = e.instantiateRev (subst.extract start stop) 247 | 248 | def abstract1 (v : FVarId) : Expr → (k :_:= 0) → Expr 249 | | .bvar i, d => .bvar (if i < d then i else i + 1) 250 | | e@(.fvar v'), d => if v == v' then .bvar d else e 251 | | .mdata m e, d => .mdata m (abstract1 v e d) 252 | | .proj s i e, d => .proj s i (abstract1 v e d) 253 | | .app f a, d => .app (abstract1 v f d) (abstract1 v a d) 254 | | .lam n t b bi, d => .lam n (abstract1 v t d) (abstract1 v b (d+1)) bi 255 | | .forallE n t b bi, d => .forallE n (abstract1 v t d) (abstract1 v b (d+1)) bi 256 | | .letE n t val b bi, d => 257 | .letE n (abstract1 v t d) (abstract1 v val d) (abstract1 v b (d+1)) bi 258 | | e@(.const ..), _ 259 | | e@(.sort _), _ 260 | | e@(.mvar _), _ 261 | | e@(.lit _), _ => e 262 | 263 | @[simp] def abstractList : Expr → List FVarId → (k :_:= 0) → Expr 264 | | e, [], _ => e 265 | | e, a :: as, k => abstractList (abstract1 a e k) as k 266 | 267 | /-- This could be an `@[implemented_by]` -/ 268 | @[simp] axiom abstract_eq (e : Expr) (xs : List FVarId) : 269 | e.abstract ⟨xs.map .fvar⟩ = e.abstractList xs 270 | 271 | /-- This could be an `@[implemented_by]` -/ 272 | @[simp] axiom abstractRange_eq (e : Expr) (n : Nat) (xs : Array Expr) : 273 | e.abstractRange n xs = e.abstract (xs.extract 0 n) 274 | 275 | def hasLooseBVar' : (e : @& Expr) → (bvarIdx : @& Nat) → Bool 276 | | .bvar i, d => i = d 277 | | .mdata _ e, d 278 | | .proj _ _ e, d => hasLooseBVar' e d 279 | | .app f a, d => hasLooseBVar' f d || hasLooseBVar' a d 280 | | .lam _ t b _, d 281 | | .forallE _ t b _, d => hasLooseBVar' t d || hasLooseBVar' b (d+1) 282 | | .letE _ t v b _, d => hasLooseBVar' t d || hasLooseBVar' v d || hasLooseBVar' b (d+1) 283 | | .const .., _ 284 | | .sort _, _ 285 | | .fvar _, _ 286 | | .mvar _, _ 287 | | .lit _, _ => false 288 | 289 | /-- This could be an `@[implemented_by]` -/ 290 | @[simp] axiom hasLooseBVar_eq (e : Expr) (n : Nat) : e.hasLooseBVar n = e.hasLooseBVar' n 291 | 292 | end Expr 293 | 294 | namespace Level 295 | 296 | def hasParam' : Level → Bool 297 | | .param .. => true 298 | | .zero | .mvar .. => false 299 | | .succ l => l.hasParam' 300 | | .max l₁ l₂ | .imax l₁ l₂ => l₁.hasParam' || l₂.hasParam' 301 | 302 | /-- This is currently false, see bug lean4#8554 -/ 303 | @[simp] axiom hasParam_eq (l : Level) : l.hasParam = l.hasParam' 304 | 305 | def hasMVar' : Level → Bool 306 | | .mvar .. => true 307 | | .zero | .param .. => false 308 | | .succ l => l.hasMVar' 309 | | .max l₁ l₂ | .imax l₁ l₂ => l₁.hasMVar' || l₂.hasMVar' 310 | 311 | /-- This is currently false, see bug lean4#8554 -/ 312 | @[simp] axiom hasMVar_eq (l : Level) : l.hasMVar = l.hasMVar' 313 | 314 | /-- This is because the `BEq` instance is implemented in C++ -/ 315 | @[instance] axiom instLawfulBEqLevel : LawfulBEq Level 316 | 317 | end Level 318 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Environment.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 | def isAsSafeAs : DefinitionSafety → ConstantInfo → Bool 9 | | .unsafe, _ => true 10 | | .partial, ci => !ci.isUnsafe 11 | | .safe, ci => !ci.isUnsafe && !ci.isPartial 12 | 13 | variable (safety : DefinitionSafety) (env : VEnv) in 14 | def TrConstant (ci : ConstantInfo) (ci' : VConstant) : Prop := 15 | isAsSafeAs safety ci ∧ ci.levelParams.length = ci'.uvars ∧ 16 | TrExprS env ci.levelParams [] ci.type ci'.type 17 | 18 | variable (safety : DefinitionSafety) (env : VEnv) in 19 | def TrConstVal (ci : ConstantInfo) (ci' : VConstVal) : Prop := 20 | TrConstant safety env ci ci'.toVConstant ∧ ci.name = ci'.name 21 | 22 | variable (safety : DefinitionSafety) (env : VEnv) in 23 | def TrDefVal (ci : ConstantInfo) (ci' : VDefVal) : Prop := 24 | TrConstVal safety env ci ci'.toVConstVal ∧ 25 | TrExprS env ci.levelParams [] ci.value! ci'.value 26 | 27 | def AddQuot1 (name : Name) (kind : QuotKind) (ci' : VConstant) (P : ConstMap → VEnv → Prop) 28 | (m : ConstMap) (env : VEnv) : Prop := 29 | ∃ levelParams type env', 30 | let ci := .quotInfo { name, kind, levelParams, type } 31 | TrConstant .safe env ci ci' ∧ 32 | env.addConst name ci' = some env' ∧ 33 | P (m.insert name ci) env' 34 | 35 | theorem AddQuot1.to_addQuot 36 | (H1 : ∀ m env, P m env → f env = some env') 37 | (m env) (H : AddQuot1 name kind ci' P m env) : 38 | env.addConst name (some ci') >>= f = some env' := by 39 | let ⟨_, _, _, h1, h2, h3⟩ := H 40 | simpa using ⟨_, h2, H1 _ _ h3⟩ 41 | 42 | def AddQuot (m₁ m₂ : ConstMap) (env₁ env₂ : VEnv) : Prop := 43 | AddQuot1 ``Quot .type quotConst (m := m₁) (env := env₁) <| 44 | AddQuot1 ``Quot.mk .ctor quotMkConst <| 45 | AddQuot1 ``Quot.lift .lift quotLiftConst <| 46 | AddQuot1 ``Quot.ind .ind quotIndConst (· = m₂ ∧ ·.addDefEq quotDefEq = env₂) 47 | 48 | nonrec theorem AddQuot.to_addQuot (H : AddQuot m₁ m₂ env₁ env₂) : env₁.addQuot = some env₂ := 49 | open AddQuot1 in (to_addQuot <| to_addQuot <| to_addQuot <| to_addQuot (by simp)) _ _ H 50 | 51 | inductive AddInduct (m₁ : ConstMap) (env₁ : VEnv) (decl : VInductDecl) 52 | (m₂ : ConstMap) (env₂ : VEnv) : Prop 53 | -- TODO 54 | 55 | nonrec theorem AddInduct.to_addInduct 56 | (H : AddInduct m₁ env₁ decl m₂ env₂) : env₁.addInduct decl = some env₂ := 57 | nomatch H 58 | 59 | variable (safety : DefinitionSafety) in 60 | inductive TrEnv' : ConstMap → Bool → VEnv → Prop where 61 | | empty : TrEnv' {} false .empty 62 | | block : 63 | ¬isAsSafeAs safety ci → 64 | env.addConst ci.name none = some env' → 65 | TrEnv' C Q env → 66 | TrEnv' (C.insert ci.name ci) Q env' 67 | | axiom : 68 | TrConstant safety env (.axiomInfo ci) ci' → 69 | ci'.WF env → 70 | env.addConst ci.name (some ci') = some env' → 71 | TrEnv' C Q env → 72 | TrEnv' (C.insert ci.name (.axiomInfo ci)) Q env' 73 | | defn {ci' : VDefVal} : 74 | TrDefVal safety env (.defnInfo ci) ci' → 75 | ci'.WF env → 76 | env.addConst ci.name (some ci'.toVConstant) = some env' → 77 | TrEnv' C Q env → 78 | TrEnv' (C.insert ci.name (.defnInfo ci)) Q (env'.addDefEq ci'.toDefEq) 79 | | opaque {ci' : VDefVal} : 80 | TrDefVal safety env (.opaqueInfo ci) ci' → 81 | ci'.WF env → 82 | env.addConst ci.name (some ci'.toVConstant) = some env' → 83 | TrEnv' C Q env → 84 | TrEnv' (C.insert ci.name (.opaqueInfo ci)) Q env' 85 | | quot : 86 | env.QuotReady → 87 | AddQuot C C' env env' → 88 | TrEnv' C false env → 89 | TrEnv' C' true env' 90 | | induct : 91 | decl.WF env → 92 | AddInduct C env decl C' env' → 93 | TrEnv' C Q env → 94 | TrEnv' C' Q env' 95 | 96 | def TrEnv (safety : DefinitionSafety) (env : Environment) (venv : VEnv) : Prop := 97 | TrEnv' safety env.constants env.quotInit venv 98 | 99 | theorem TrEnv'.wf (H : TrEnv' safety C Q venv) : venv.WF := by 100 | induction H with 101 | | empty => exact ⟨_, .empty⟩ 102 | | block _ h _ ih => 103 | have ⟨_, H⟩ := ih 104 | exact ⟨_, H.decl <| .block h⟩ 105 | | «axiom» _ h1 h2 _ ih => 106 | have ⟨_, H⟩ := ih 107 | exact ⟨_, H.decl <| .axiom (ci := ⟨_, _⟩) h1 h2⟩ 108 | | defn h1 h2 h3 _ ih => 109 | have ⟨_, H⟩ := ih 110 | have := h1.1.2; dsimp [ConstantInfo.name, ConstantInfo.toConstantVal] at this 111 | exact ⟨_, H.decl <| .def h2 (this ▸ h3)⟩ 112 | | «opaque» h1 h2 h3 _ ih => 113 | have ⟨_, H⟩ := ih 114 | have := h1.1.2; dsimp [ConstantInfo.name, ConstantInfo.toConstantVal] at this 115 | exact ⟨_, H.decl <| .opaque h2 (this ▸ h3)⟩ 116 | | quot h1 h2 _ ih => 117 | have ⟨_, H⟩ := ih 118 | exact ⟨_, H.decl <| .quot h1 h2.to_addQuot⟩ 119 | | induct h1 h2 _ ih => 120 | have ⟨_, H⟩ := ih 121 | exact ⟨_, H.decl <| .induct h1 h2.to_addInduct⟩ 122 | -------------------------------------------------------------------------------- /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, 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, 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, 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).snd = none → s = none ∧ _ from 96 | (this _ · |>.2) 97 | have {l} (hmv : l.hasMVar' = false) 98 | {g} (H : ∀ {s'}, (g.run s').snd = none → s' = none ∧ 99 | ((getUndefParam.F Us l).run none = (true, none) → ∃ u', VLevel.ofLevel Us l = some u')) (s) : 100 | ((do if (!(← getUndefParam.F Us l)) = true then pure PUnit.unit else g).run s).snd = none → 101 | s = none ∧ ∃ u', VLevel.ofLevel Us l = some u' := by 102 | simp; split <;> rename_i h 103 | · simp; revert h 104 | simp [getUndefParam.F]; split <;> [simp; split <;> [split <;> simp; simp]] 105 | rintro rfl; simp at * 106 | exact ofLevel_of_not_hasParam Us ‹_› hmv 107 | · refine fun h' => let ⟨h1, h2⟩ := H h'; have := ?_; ⟨this, h2 ?_⟩ 108 | · revert h h1 109 | simp [getUndefParam.F]; split <;> [simp; split <;> [split <;> simp; simp]] 110 | · revert h h1; subst s 111 | cases (getUndefParam.F Us l).run none; simp; rintro rfl rfl; rfl 112 | have lt {n a} : n + 1 < a → n < a := by omega 113 | induction l with ( 114 | refine this hmv fun h => ?_; clear this 115 | simp [hasMVar', VLevel.ofLevel, *] at *) 116 | | succ _ ih => 117 | have ⟨h, _, h1⟩ := ih hmv _ h 118 | exact ⟨h, fun _ => ⟨_, _, h1, rfl⟩⟩ 119 | | max _ _ ih1 ih2 | imax _ _ ih1 ih2 => 120 | have ⟨h, _, h2⟩ := ih2 hmv.2 _ h 121 | have ⟨h, _, h1⟩ := ih1 hmv.1 _ h 122 | exact ⟨h, fun _ => ⟨_, _, h1, _, h2, rfl⟩⟩ 123 | | param => 124 | simp [getUndefParam.F, Id, hasParam', List.idxOf_lt_length_iff, *] 125 | split <;> simp [*] 126 | | _ => simp [*] 127 | 128 | variable (s : Name → Level) in 129 | def substParams' (red : Bool) : Level → Level 130 | | .zero => .zero 131 | | .succ v => .succ (substParams' (v.hasParam ∧ red) v) 132 | | .max v₁ v₂ => 133 | let red := (v₁.hasParam ∨ v₂.hasParam) ∧ red 134 | (if red then mkLevelMax' else .max) (substParams' red v₁) (substParams' red v₂) 135 | | .imax v₁ v₂ => 136 | let red := (v₁.hasParam ∨ v₂.hasParam) ∧ red 137 | (if red then mkLevelIMax' else .imax) (substParams' red v₁) (substParams' red v₂) 138 | | .param n => s n 139 | | u => u 140 | 141 | theorem substParams_eq_self {u : Level} (h : u.hasParam' = false) : 142 | substParams' s red u = u := by 143 | induction u generalizing red <;> simp_all [substParams', hasParam'] 144 | 145 | @[simp] theorem substParams_eq (u : Level) (s : Name → Option Level) : 146 | substParams u s = substParams' (fun x => (s x).getD (.param x)) true u := by 147 | unfold substParams 148 | induction u <;> simp [substParams.go, substParams', hasParam', ← Bool.or_eq_true] <;> 149 | split <;> simp [*, substParams_eq_self] <;> simp_all [substParams_eq_self] 150 | 151 | theorem substParams_id {u : Level} : 152 | substParams' .param false u = u := by induction u <;> simp_all [substParams', hasParam'] 153 | -------------------------------------------------------------------------------- /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, Nat.sub_zero, List.drop_zero, 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 [← List.map_reverse, 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 [Expr.abstractList_append, Expr.abstract1, 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 | variable (env : VEnv) (Us : List Name) in 201 | inductive TrLCtx' : List LocalDecl → VLCtx → Prop 202 | | nil : TrLCtx' [] [] 203 | | cons : 204 | TrLCtx' ds Δ → TrLocalDecl env Us Δ d d' → 205 | TrLCtx' (d :: ds) ((some d.fvarId, d') :: Δ) 206 | 207 | def TrLCtx (env : VEnv) (Us : List Name) (lctx : LocalContext) (Δ : VLCtx) : Prop := 208 | lctx.WF ∧ TrLCtx' env Us lctx.toList Δ 209 | 210 | theorem TrLCtx'.noBV : TrLCtx' env Us ds Δ → Δ.NoBV 211 | | .nil => rfl 212 | | .cons h _ => h.noBV 213 | 214 | theorem TrLCtx'.forall₂ : 215 | TrLCtx' env Us ds Δ → ds.Forall₂ Δ (R := fun d d' => d'.1 = some d.fvarId) 216 | | .nil => by simp [LocalContext.toList] 217 | | .cons h _ => by simp [LocalContext.toList]; exact h.forall₂ 218 | 219 | theorem TrLCtx'.fvars_eq (H : TrLCtx' env Us ds Δ) : ds.map (·.fvarId) = Δ.fvars := by 220 | simp [VLCtx.fvars] 221 | induction H with 222 | | nil => rfl 223 | | cons h1 _ ih => simp [h1, ← ih] 224 | 225 | theorem TrLCtx.fvars_eq (H : TrLCtx env Us lctx Δ) : lctx.fvars = Δ.fvars := 226 | H.2.fvars_eq 227 | 228 | theorem TrLCtx'.find?_isSome (H : TrLCtx' env Us ds Δ) : 229 | (ds.find? (fv == ·.fvarId)).isSome = (Δ.find? (.inr fv)).isSome := by 230 | rw [← VLCtx.lookup_isSome] 231 | induction H with 232 | | nil => rfl 233 | | @cons _ _ d d' _ ih => 234 | simp [List.find?, List.lookup, show (some fv == some d.fvarId) = (fv == d.fvarId) from rfl] 235 | split <;> simp [*] 236 | 237 | theorem TrLCtx.find?_isSome (H : TrLCtx env Us lctx Δ) : 238 | (lctx.find? fv).isSome = (Δ.find? (.inr fv)).isSome := by 239 | rw [H.1.find?_eq_find?_toList, H.2.find?_isSome] 240 | 241 | theorem TrLCtx'.find?_eq_some (H : TrLCtx' env Us ds Δ) : 242 | (∃ d, ds.find? (fv == ·.fvarId) = some d) ↔ fv ∈ Δ.fvars := by 243 | rw [← Option.isSome_iff_exists, H.find?_isSome, Option.isSome_iff_exists, VLCtx.find?_eq_some] 244 | 245 | theorem TrLCtx.find?_eq_some (H : TrLCtx env Us lctx Δ) : 246 | (∃ d, lctx.find? fv = some d) ↔ fv ∈ Δ.fvars := by 247 | rw [H.1.find?_eq_find?_toList, H.2.find?_eq_some] 248 | 249 | theorem TrLCtx.find?_eq_none (H : TrLCtx env Us lctx Δ) : 250 | lctx.find? fv = none ↔ ¬fv ∈ Δ.fvars := by simp [← H.find?_eq_some] 251 | 252 | theorem TrLCtx.contains (H : TrLCtx env Us lctx Δ) : lctx.contains fv ↔ fv ∈ Δ.fvars := by 253 | rw [LocalContext.contains, PersistentHashMap.find?_isSome, Option.isSome_iff_exists] 254 | exact H.find?_eq_some 255 | 256 | theorem TrLCtx'.wf : TrLCtx' env Us ds Δ → (ds.map (·.fvarId)).Nodup → Δ.WF env Us.length 257 | | .nil, _ => ⟨⟩ 258 | | .cons h1 h2, .cons H1 H2 => 259 | ⟨h1.wf H2, by simpa [← h1.find?_eq_some] using H1, h2.wf⟩ 260 | 261 | theorem TrLCtx.wf (H : TrLCtx env Us lctx Δ) : Δ.WF env Us.length := H.2.wf H.1.nodup 262 | 263 | theorem TrLCtx'.find?_of_mem (henv : env.WF) (H : TrLCtx' env Us ds Δ) 264 | (nd : (ds.map (·.fvarId)).Nodup) (hm : decl ∈ ds) : 265 | ∃ e A, Δ.find? (.inr decl.fvarId) = some (e, A) ∧ TrExprS env Us Δ decl.type A := by 266 | have := H.wf nd 267 | match H with 268 | | .nil => cases hm 269 | | .cons (ds := ds) h1 h2 => 270 | simp [VLCtx.find?, VLCtx.next] 271 | obtain _ | ⟨_, hm : decl ∈ ds⟩ := hm 272 | · simp [and_assoc] 273 | cases h2 with 274 | | vlam h2 h3 => exact h2.weakFV henv (.skip_fvar _ _ .refl) this 275 | | vlet h2 h3 => simpa [VLocalDecl.depth] using h2.weakFV henv (.skip_fvar _ _ .refl) this 276 | · simp at nd; rw [if_neg (by simpa using Ne.symm (nd.1 _ hm))]; simp 277 | have ⟨_, _, h1, h2⟩ := h1.find?_of_mem henv nd.2 hm 278 | refine ⟨_, _, ⟨_, _, h1, rfl, rfl⟩, ?_⟩ 279 | simpa using h2.weakFV henv (.skip_fvar _ _ .refl) this 280 | 281 | theorem TrLCtx.find?_of_mem (henv : env.WF) (H : TrLCtx env Us lctx Δ) 282 | (hm : decl ∈ lctx.toList) : 283 | ∃ e A, Δ.find? (.inr decl.fvarId) = some (e, A) ∧ TrExprS env Us Δ decl.type A := 284 | H.2.find?_of_mem henv H.1.nodup hm 285 | 286 | theorem TrLCtx.mkLocalDecl 287 | (h1 : TrLCtx env Us lctx Δ) (h2 : lctx.find? fv = none) (h3 : TrExprS env Us Δ ty ty') 288 | (h4 : env.IsType Us.length Δ.toCtx ty') : 289 | TrLCtx env Us (lctx.mkLocalDecl fv name ty bi kind) ((some fv, .vlam ty') :: Δ) := 290 | ⟨h1.1.mkLocalDecl h2, by simpa using .cons h1.2 (.vlam h3 h4)⟩ 291 | 292 | theorem TrLCtx.mkLetDecl 293 | (h1 : TrLCtx env Us lctx Δ) (h2 : lctx.find? fv = none) 294 | (h3 : TrExprS env Us Δ ty ty') (h4 : TrExprS env Us Δ val val') 295 | (h5 : env.HasType Us.length Δ.toCtx val' ty') : 296 | TrLCtx env Us (lctx.mkLetDecl fv name ty val bi kind) ((some fv, .vlet ty' val') :: Δ) := 297 | ⟨h1.1.mkLetDecl h2, by simpa using .cons h1.2 (.vlet h3 h4 h5)⟩ 298 | -------------------------------------------------------------------------------- /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/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.mono (H : ngen₁ ≤ ngen₂) : 11 | ConditionallyTyped ngen₁ env Us Δ e → ConditionallyTyped ngen₂ env Us Δ e 12 | | ⟨h1, h2, h3⟩ => ⟨h1, h2.mono fun _ h => h.mono H, h3⟩ 13 | 14 | theorem ConditionallyTyped.weakN_inv 15 | (henv : VEnv.WF env) (hΔ : VLCtx.WF env Us.length ((some fv, d) :: Δ)) 16 | (H : ConditionallyTyped ngen env Us ((some fv, d) :: Δ) e) : 17 | ConditionallyTyped ngen env Us Δ e := by 18 | refine ⟨H.1, H.2.1, fun H2 => ?_⟩ 19 | have := H.2.2 20 | have ⟨e', h⟩ := H.2.2 H2.fvars_cons 21 | exact TrExprS.weakFV_inv henv (.skip_fvar _ _ .refl) (.refl henv hΔ) h H.1 H2 22 | 23 | theorem ConditionallyTyped.fresh 24 | (henv : Ordered env) (hΔ : VLCtx.WF env Us.length ((some ⟨ngen.curr⟩, d) :: Δ)) 25 | (H : ConditionallyTyped ngen env Us Δ e) : 26 | ConditionallyTyped ngen env Us ((some ⟨ngen.curr⟩, d) :: Δ) e := by 27 | refine have ⟨H1, H2, H3⟩ := H; ⟨H1, H2, fun H4 => ?_⟩ 28 | refine have ⟨_, h⟩ := H3 (H4.mp ?_ H2); ⟨_, h.weakFV henv (.skip_fvar _ _ .refl) hΔ⟩ 29 | intro _ h1 h2; simp at h1; rcases h1 with rfl | h1 30 | · cases Nat.lt_irrefl _ (h2 _ rfl) 31 | · exact h1 32 | 33 | def TrHasType (env : VEnv) (Us : List Name) (Δ : VLCtx) (e A : Expr) : Prop := 34 | (∀ P, IsFVarUpSet P Δ → FVarsIn P e → FVarsIn P A) ∧ 35 | ∃ e' A', TrExprS env Us Δ e e' ∧ TrExprS env Us Δ A A' ∧ env.HasType Us.length Δ.toCtx e' A' 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 → TrHasType env Us Δ e A) 41 | 42 | theorem ConditionallyHasType.typed : 43 | ConditionallyHasType ngen env Us Δ e A → ConditionallyTyped ngen env Us Δ e 44 | | ⟨c1, f1, _, _, H⟩ => ⟨c1, f1, fun h => let ⟨_, _, _, h, _⟩ := H h; ⟨_, h⟩⟩ 45 | 46 | theorem ConditionallyHasType.mono (H : ngen₁ ≤ ngen₂) : 47 | ConditionallyHasType ngen₁ env Us Δ e A → ConditionallyHasType ngen₂ env Us Δ e A 48 | | ⟨c1, f1, c2, f2, h'⟩ => ⟨c1, f1.mono fun _ h => h.mono H, c2, f2.mono fun _ h => h.mono H, h'⟩ 49 | 50 | theorem ConditionallyHasType.weakN_inv 51 | (henv : VEnv.WF env) (hΔ : VLCtx.WF env Us.length ((some fv, d) :: Δ)) 52 | (H : ConditionallyHasType ngen env Us ((some fv, d) :: Δ) e A) : 53 | ConditionallyHasType ngen env Us Δ e A := by 54 | have ⟨c1, f1, c2, f2, H⟩ := H 55 | refine ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 56 | have ⟨h1, e', A', h2, h3, h4⟩ := H H4.fvars_cons 57 | have W : VLCtx.FVLift Δ ((some fv, d) :: Δ) 0 (0 + d.depth) 0 := .skip_fvar _ _ .refl 58 | have ⟨e'', he⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h2 c1 H4 59 | have ee := h2.uniq henv (.refl henv hΔ) <| he.weakFV henv W hΔ 60 | have : IsFVarUpSet (· ∈ VLCtx.fvars Δ) ((some fv, d) :: Δ) := ⟨(hΔ.2.1 _ rfl).elim, .fvars⟩ 61 | have ⟨_, hA⟩ := TrExprS.weakFV_inv henv W (.refl henv hΔ) h3 c2 <| h1 _ this H4 62 | have AA := h3.uniq henv (.refl henv hΔ) <| hA.weakFV henv W hΔ 63 | have h4 := h4.defeqU_r henv hΔ.toCtx AA |>.defeqU_l henv hΔ.toCtx ee 64 | have h4 := (HasType.weakN_iff henv hΔ.toCtx W.toCtx).1 h4 65 | refine ⟨fun P hP he' => ?_, _, _, he, hA, h4⟩ 66 | exact h1 _ 67 | ⟨fun h => (hΔ.2.1 _ rfl).elim h.2, IsFVarUpSet.and_fvars.1 hP⟩ 68 | (he'.mp (fun _ => .intro) he.fvarsIn) |>.mono fun _ => (·.1) 69 | 70 | theorem ConditionallyHasType.fresh 71 | (henv : Ordered env) (hΔ : VLCtx.WF env Us.length ((some ⟨ngen.curr⟩, d) :: Δ)) 72 | (H : ConditionallyHasType ngen env Us Δ e A) : 73 | ConditionallyHasType ngen env Us ((some ⟨ngen.curr⟩, d) :: Δ) e A := by 74 | refine have ⟨c1, f1, c2, f2, H⟩ := H; ⟨c1, f1, c2, f2, fun H4 => ?_⟩ 75 | have ⟨h1, _, _, h2, h3, h4⟩ := H (H4.mp ?_ f1) 76 | · have W : VLCtx.FVLift Δ ((some ⟨ngen.curr⟩, d) :: Δ) 0 (0 + d.depth) 0 := .skip_fvar _ _ .refl 77 | refine ⟨fun P hP => h1 _ hP.2, ?_⟩ 78 | exact ⟨_, _, h2.weakFV henv W hΔ, h3.weakFV henv W hΔ, h4.weakN henv W.toCtx⟩ 79 | · intro _ h1 h2; simp at h1; rcases h1 with rfl | h1 80 | · cases Nat.lt_irrefl _ (h2 _ rfl) 81 | · exact h1 82 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Typing/EnvLemmas.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Std.SMap 2 | import Lean4Lean.Verify.Environment 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 | section 21 | 22 | variable (safety : DefinitionSafety) (env : VEnv) in 23 | inductive Aligned1 : ConstantInfo → Option VConstant → Prop where 24 | | block : ¬isAsSafeAs safety ci → Aligned1 ci none 25 | | const : TrConstant safety env ci ci' → Aligned1 ci (some ci') 26 | 27 | variable (safety : DefinitionSafety) in 28 | inductive Aligned : ConstMap → VEnv → Prop where 29 | | empty : Aligned {} .empty 30 | | const : Aligned C venv → Aligned1 safety venv ci ci' → 31 | venv.addConst n ci' = some venv' → Aligned (C.insert n ci) venv' 32 | | defeq : Aligned C venv → Aligned C (venv.addDefEq df) 33 | 34 | theorem Aligned.map_wf' (H : Aligned safety C venv) : 35 | C.WF ∧ ∀ n, (C.find? n).isSome = (venv.constants n).isSome := by 36 | induction H with 37 | | empty => exact ⟨.empty, by simp [SMap.find?, VEnv.empty]⟩ 38 | | @const _ _ _ _ n _ _ _ eq ih => 39 | simp [VEnv.addConst] at eq; split at eq <;> cases eq 40 | have := ih.2 n; simp_all [-ih] 41 | refine ⟨.insert ih.1 _ _ this, fun n' => ?_⟩ 42 | simp [ih.1.find?_insert]; split <;> simp [ih.2] 43 | | defeq _ ih => exact ih 44 | 45 | theorem Aligned.map_wf (H : Aligned safety C venv) : C.WF := H.map_wf'.1 46 | theorem Aligned.find?_isSome (H : Aligned safety C venv) : 47 | (C.find? n).isSome = (venv.constants n).isSome := H.map_wf'.2 _ 48 | 49 | theorem Aligned.addQuot1 {Q : Prop} 50 | (H1 : ∀ c env, Aligned safety c env → P c env → Q) 51 | (C env) (wf : Aligned safety C env) (H2 : AddQuot1 n k ci P C env) : Q := by 52 | let ⟨_, _, h1, h2, h3, h4⟩ := H2 53 | refine H1 _ _ (wf.const (.const ⟨by cases safety <;> rfl, h2.2⟩) h3) h4 54 | 55 | nonrec theorem Aligned.addQuot (H : AddQuot C₁ C₂ venv₁ venv₂) 56 | (wf : Aligned safety C₁ venv₁) : Aligned safety C₂ venv₂ := by 57 | dsimp [AddQuot] at H 58 | refine (addQuot1 <| addQuot1 <| addQuot1 <| addQuot1 ?_) _ _ wf H 59 | rintro _ _ h ⟨rfl, rfl⟩; exact h.defeq 60 | 61 | theorem Aligned.addInduct (H : AddInduct C₁ venv₁ decl C₂ venv₂) : 62 | Aligned safety C₁ env₁ → Aligned safety C₂ env₂ := 63 | nomatch H 64 | 65 | theorem TrEnv'.aligned (H : TrEnv' safety C Q venv) : Aligned safety C venv := by 66 | induction H with 67 | | empty => exact .empty 68 | | block hs h _ ih => exact ih.const (.block hs) h 69 | | «axiom» h1 _ h _ ih => exact ih.const (.const h1) h 70 | | «opaque» h1 _ h _ ih => exact ih.const (.const h1.1.1) h 71 | | defn h1 _ h _ ih => exact (ih.const (.const h1.1.1) h).defeq 72 | | quot _ h _ ih => exact ih.addQuot h 73 | | induct _ h _ ih => exact ih.addInduct h 74 | 75 | theorem TrEnv'.map_wf (H : TrEnv' safety C Q venv) : C.WF := H.aligned.map_wf 76 | 77 | theorem Aligned.find? (H : Aligned safety C venv) 78 | (h : C.find? name = some ci) (hs : isAsSafeAs safety ci) : 79 | ∃ ci', venv.constants name = some (some ci') ∧ TrConstant safety venv ci ci' := by 80 | have mono {env₁ env₂} (H : env₁.LE env₂) : 81 | (∃ ci', env₁.constants name = some (some ci') ∧ TrConstant safety env₁ ci ci') → 82 | (∃ ci', env₂.constants name = some (some ci') ∧ TrConstant safety env₂ ci ci') 83 | | ⟨_, h1, h2⟩ => ⟨_, H.constants h1, h2.mono H⟩ 84 | induction H with 85 | | empty => simp [SMap.find?] at h 86 | | const h1 h2 h3 ih => 87 | have := VEnv.addConst_le h3 88 | rw [h1.map_wf.find?_insert] at h; split at h 89 | · cases h; rename_i n _ h; simp at h; subst n 90 | simp [VEnv.addConst] at h3; split at h3 <;> cases h3 91 | cases h2; · contradiction 92 | simp; rename_i h'; exact h'.mono this 93 | · let ⟨_, h1, h2⟩ := ih h; exact ⟨_, this.constants h1, h2.mono this⟩ 94 | | defeq h1 ih => let ⟨_, h1, h2⟩ := ih h; exact ⟨_, h1, h2.mono VEnv.addDefEq_le⟩ 95 | 96 | theorem Aligned.find?_uniq (H : Aligned safety C venv) 97 | (h : C.find? name = some ci) (hs : venv.constants name = some (some ci')) : 98 | TrConstant safety venv ci ci' := by 99 | induction H with 100 | | empty => simp [SMap.find?] at h 101 | | const h1 h2 h3 ih => 102 | have := VEnv.addConst_le h3 103 | simp [VEnv.addConst] at h3; split at h3 <;> cases h3 104 | simp [h1.map_wf.find?_insert] at h hs; revert h hs; split 105 | · rintro ⟨⟩ ⟨⟩; rename_i n _ _ _; subst n 106 | let .const h2 := h2; exact h2.mono this 107 | · intro hs h; exact (ih h hs).mono this 108 | | defeq h1 ih => exact (ih h hs).mono VEnv.addDefEq_le 109 | 110 | end 111 | 112 | theorem TrEnv.find? (H : TrEnv safety env venv) 113 | (h : env.find? name = some ci) (hs : isAsSafeAs safety ci) : 114 | ∃ ci', venv.constants name = some (some ci') ∧ TrConstant safety venv ci ci' := 115 | H.aligned.find? (H.map_wf.find?'_eq_find? _ ▸ h) hs 116 | 117 | theorem TrEnv.find?_uniq (H : TrEnv safety env venv) 118 | (h : env.find? name = some ci) (hs : venv.constants name = some (some ci')) : 119 | TrConstant safety venv ci ci' := 120 | H.aligned.find?_uniq (H.map_wf.find?'_eq_find? _ ▸ h) hs 121 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/Typing/Expr.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.Theory.Typing.Basic 2 | import Lean4Lean.Verify.NameGenerator 3 | import Lean4Lean.Verify.VLCtx 4 | 5 | namespace Lean4Lean 6 | open Lean 7 | 8 | def Closed : Expr → (k :_:= 0) → Prop 9 | | .bvar i, k => i < k 10 | | .fvar _, _ | .sort .., _ | .const .., _ | .lit .., _ => True 11 | | .app f a, k => Closed f k ∧ Closed a k 12 | | .lam _ d b _, k 13 | | .forallE _ d b _, k => Closed d k ∧ Closed b (k+1) 14 | | .letE _ d v b _, k => Closed d k ∧ Closed v k ∧ Closed b (k+1) 15 | | .proj _ _ e, k | .mdata _ e, k => Closed e k 16 | | .mvar .., _ => False 17 | 18 | nonrec abbrev _root_.Lean.Expr.Closed := @Closed 19 | 20 | variable (fvars : FVarId → Prop) in 21 | def FVarsIn : Expr → Prop 22 | | .bvar _ => True 23 | | .fvar fv => fvars fv 24 | | .sort .. | .const .. | .lit .. => True 25 | | .app f a => FVarsIn f ∧ FVarsIn a 26 | | .lam _ d b _ 27 | | .forallE _ d b _ => FVarsIn d ∧ FVarsIn b 28 | | .letE _ d v b _ => FVarsIn d ∧ FVarsIn v ∧ FVarsIn b 29 | | .proj _ _ e | .mdata _ e => FVarsIn e 30 | | .mvar .. => False 31 | 32 | nonrec abbrev _root_.Lean.Expr.FVarsIn := @FVarsIn 33 | 34 | def VLocalDecl.WF (env : VEnv) (U : Nat) (Γ : List VExpr) : VLocalDecl → Prop 35 | | .vlam type => env.IsType U Γ type 36 | | .vlet type value => env.HasType U Γ value type 37 | 38 | variable (env : VEnv) (U : Nat) in 39 | def VLCtx.WF : VLCtx → Prop 40 | | [] => True 41 | | (ofv, d) :: (Δ : VLCtx) => 42 | VLCtx.WF Δ ∧ (∀ fv, ofv = some fv → fv ∉ Δ.fvars) ∧ VLocalDecl.WF env U Δ.toCtx d 43 | 44 | def TrProj : ∀ (Γ : List VExpr) (structName : Name) (idx : Nat) (e : VExpr), VExpr → Prop := sorry 45 | 46 | variable (env : VEnv) (Us : List Name) in 47 | inductive TrExprS : VLCtx → Expr → VExpr → Prop 48 | | bvar : Δ.find? (.inl i) = some (e, A) → TrExprS Δ (.bvar i) e 49 | | fvar : Δ.find? (.inr fv) = some (e, A) → TrExprS Δ (.fvar fv) e 50 | | sort : VLevel.ofLevel Us u = some u' → TrExprS Δ (.sort u) (.sort u') 51 | | const : 52 | env.constants c = some (some ci) → 53 | us.mapM (VLevel.ofLevel Us) = some us' → 54 | us.length = ci.uvars → 55 | TrExprS Δ (.const c us) (.const c us') 56 | | app : 57 | env.HasType Us.length Δ.toCtx f' (.forallE A B) → 58 | env.HasType Us.length Δ.toCtx a' A → 59 | TrExprS Δ f f' → TrExprS Δ a a' → TrExprS Δ (.app f a) (.app f' a') 60 | | lam : 61 | env.IsType Us.length Δ.toCtx ty' → 62 | TrExprS Δ ty ty' → TrExprS ((none, .vlam ty') :: Δ) body body' → 63 | TrExprS Δ (.lam name ty body bi) (.lam ty' body') 64 | | forallE : 65 | env.IsType Us.length Δ.toCtx ty' → 66 | env.IsType Us.length (ty' :: Δ.toCtx) body' → 67 | TrExprS Δ ty ty' → TrExprS ((none, .vlam ty') :: Δ) body body' → 68 | TrExprS Δ (.forallE name ty body bi) (.forallE ty' body') 69 | | letE : 70 | env.HasType Us.length Δ.toCtx val' ty' → 71 | TrExprS Δ ty ty' → TrExprS Δ val val' → 72 | TrExprS ((none, .vlet ty' val') :: Δ) body body' → 73 | TrExprS Δ (.letE name ty val body nd) body' 74 | | lit : TrExprS Δ l.toConstructor e → TrExprS Δ (.lit l) e 75 | | mdata : TrExprS Δ e e' → TrExprS Δ (.mdata d e) e' 76 | | proj : TrExprS Δ e e' → TrProj Δ.toCtx s i e' e'' → TrExprS Δ (.proj s i e) e'' 77 | 78 | def TrExpr (env : VEnv) (Us : List Name) (Δ : VLCtx) (e : Expr) (e' : VExpr) : Prop := 79 | ∃ e₂, TrExprS env Us Δ e e₂ ∧ env.IsDefEqU Us.length Δ.toCtx e₂ e' 80 | -------------------------------------------------------------------------------- /Lean4Lean/Verify/VLCtx.lean: -------------------------------------------------------------------------------- 1 | import Lean4Lean.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.liftN : VLocalDecl → Nat → Nat → VLocalDecl 29 | | .vlam A, n, k => .vlam (A.liftN n k) 30 | | .vlet A e, n, k => .vlet (A.liftN n k) (e.liftN n k) 31 | 32 | def VLocalDecl.inst : VLocalDecl → VExpr → (k : Nat := 0) → VLocalDecl 33 | | .vlam A, e₀, k => .vlam (A.inst e₀ k) 34 | | .vlet A e, e₀, k => .vlet (A.inst e₀ k) (e.inst e₀ k) 35 | 36 | def VLocalDecl.instL : VLocalDecl → List VLevel → VLocalDecl 37 | | .vlam A, ls => .vlam (A.instL ls) 38 | | .vlet A e, ls => .vlet (A.instL ls) (e.instL ls) 39 | 40 | def VLCtx := List (Option FVarId × VLocalDecl) 41 | 42 | namespace VLCtx 43 | 44 | def bvars : VLCtx → Nat 45 | | [] => 0 46 | | (none, _) :: Δ => bvars Δ + 1 47 | | (some _, _) :: Δ => bvars Δ 48 | 49 | abbrev NoBV (Δ : VLCtx) : Prop := Δ.bvars = 0 50 | 51 | def next : Option FVarId → Nat ⊕ FVarId → Option (Nat ⊕ FVarId) 52 | | none, .inl 0 => none 53 | | none, .inl (n+1) => some (.inl n) 54 | | some _, .inl n => some (.inl n) 55 | | none, .inr fv' => some (.inr fv') 56 | | some fv, .inr fv' => if fv == fv' then none else some (.inr fv') 57 | 58 | def find? : VLCtx → Nat ⊕ FVarId → Option (VExpr × VExpr) 59 | | [], _ => none 60 | | (ofv, d) :: Δ, v => 61 | match next ofv v with 62 | | none => some (d.value, d.type) 63 | | some v => do let (e, A) ← find? Δ v; some (e.liftN d.depth, A.liftN d.depth) 64 | 65 | def liftVar (n k : Nat) : Nat ⊕ FVarId → Nat ⊕ FVarId 66 | | .inl i => .inl (if i < k then i else i + n) 67 | | .inr fv => .inr fv 68 | 69 | def varToExpr : Nat ⊕ FVarId → Expr 70 | | .inl i => .bvar i 71 | | .inr fv => .fvar fv 72 | 73 | def vlamName : VLCtx → Nat → Option (Option FVarId) 74 | | [], _ => none 75 | | (_, .vlet ..) :: Δ, i 76 | | (_, .vlam ..) :: Δ, i+1 => vlamName Δ i 77 | | (ofv, .vlam ..) :: _, 0 => some ofv 78 | 79 | def fvars (Δ : VLCtx) : List FVarId := Δ.filterMap (·.1) 80 | 81 | @[simp] theorem fvars_nil : fvars [] = [] := rfl 82 | @[simp] theorem fvars_cons_none {Δ : VLCtx} : fvars ((none, d) :: Δ) = fvars Δ := rfl 83 | @[simp] theorem fvars_cons_some {Δ : VLCtx} : fvars ((some fv, d) :: Δ) = fv :: fvars Δ := rfl 84 | 85 | def toCtx : VLCtx → List VExpr 86 | | [] => [] 87 | | (_, .vlam ty) :: Δ => ty :: VLCtx.toCtx Δ 88 | | (_, .vlet _ _) :: Δ => VLCtx.toCtx Δ 89 | 90 | def instL (Δ : VLCtx) (ls : List VLevel) : VLCtx := 91 | match Δ with 92 | | [] => [] 93 | | (ofv, d) :: Δ => (ofv, d.instL ls) :: instL Δ ls 94 | 95 | theorem lookup_isSome : ∀ {Δ : VLCtx}, (Δ.lookup (some fv)).isSome = (Δ.find? (.inr fv)).isSome 96 | | [] => rfl 97 | | (ofv, d) :: Δ => by 98 | simp [List.lookup, find?, next] 99 | cases ofv with 100 | | none => 101 | simp [show (some fv == none) = false from rfl, lookup_isSome] 102 | cases find? Δ (.inr fv) <;> simp [bind] 103 | | some fv' => 104 | simp [show (some fv == some fv') = (fv == fv') from rfl, beq_comm fv] 105 | cases fv' == fv <;> simp [lookup_isSome] 106 | cases find? Δ (.inr fv) <;> simp [bind] 107 | 108 | theorem lookup_eq_some : ∀ {Δ : VLCtx}, (∃ x, Δ.lookup (some fv) = some x) ↔ fv ∈ fvars Δ 109 | | [] => by simp 110 | | (ofv, d) :: Δ => by 111 | cases ofv with 112 | | none => simp [List.lookup, show (some fv == none) = false from rfl, lookup_eq_some] 113 | | some fv' => 114 | simp [List.lookup, show (some fv == some fv') = (fv == fv') from rfl] 115 | cases e : fv == fv' <;> simp at e <;> simp [e, lookup_eq_some] 116 | 117 | theorem find?_eq_some : (∃ x, Δ.find? (.inr fv) = some x) ↔ fv ∈ fvars Δ := by 118 | rw [← Option.isSome_iff_exists, ← lookup_isSome, Option.isSome_iff_exists, lookup_eq_some] 119 | 120 | theorem vlamName_mem_fvars : ∀ {Δ : VLCtx} {i}, Δ.vlamName i = some (some fv) → fv ∈ fvars Δ 121 | | (none, .vlet ..) :: Δ, _, h 122 | | (none, .vlam ..) :: Δ, _+1, h => vlamName_mem_fvars (Δ := Δ) h 123 | | (some _, .vlet ..) :: Δ, _, h 124 | | (some _, .vlam ..) :: Δ, _+1, h => .tail _ <| vlamName_mem_fvars (Δ := Δ) h 125 | | (some _fv, .vlam ..) :: _, 0, rfl => .head _ 126 | 127 | end VLCtx 128 | -------------------------------------------------------------------------------- /Main.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Scott Morrison. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Scott Morrison 5 | -/ 6 | import Lean.CoreM 7 | import Lean.Util.FoldConsts 8 | import Lean4Lean.Environment 9 | import Lake.Load.Manifest 10 | 11 | namespace Lean 12 | 13 | def HashMap.keyNameSet (m : Std.HashMap Name α) : NameSet := 14 | m.fold (fun s n _ => s.insert n) {} 15 | 16 | namespace Environment 17 | 18 | def importsOf (env : Environment) (n : Name) : Array Import := 19 | if n = env.header.mainModule then 20 | env.header.imports 21 | else match env.getModuleIdx? n with 22 | | .some idx => env.header.moduleData[idx.toNat]!.imports 23 | | .none => #[] 24 | 25 | end Environment 26 | 27 | namespace NameSet 28 | 29 | def ofList (names : List Name) : NameSet := 30 | names.foldl (fun s n => s.insert n) {} 31 | 32 | end NameSet 33 | 34 | /-- Like `Expr.getUsedConstants`, but produce a `NameSet`. -/ 35 | def Expr.getUsedConstants' (e : Expr) : NameSet := 36 | e.foldConsts {} fun c cs => cs.insert c 37 | 38 | namespace ConstantInfo 39 | 40 | /-- Return all names appearing in the type or value of a `ConstantInfo`. -/ 41 | def getUsedConstants (c : ConstantInfo) : NameSet := 42 | c.type.getUsedConstants' ++ match c.value? with 43 | | some v => v.getUsedConstants' 44 | | none => match c with 45 | | .inductInfo val => .ofList val.ctors 46 | | .opaqueInfo val => val.value.getUsedConstants' 47 | | .ctorInfo val => ({} : NameSet).insert val.name 48 | | .recInfo val => .ofList val.all 49 | | _ => {} 50 | 51 | end ConstantInfo 52 | 53 | end Lean 54 | 55 | open Lean hiding Environment Exception 56 | open Kernel 57 | 58 | structure Context where 59 | newConstants : Std.HashMap Name ConstantInfo 60 | verbose := false 61 | compare := false 62 | 63 | structure State where 64 | env : Environment 65 | remaining : NameSet := {} 66 | pending : NameSet := {} 67 | postponedConstructors : NameSet := {} 68 | postponedRecursors : NameSet := {} 69 | 70 | abbrev M := ReaderT Context <| StateRefT State IO 71 | 72 | /-- Check if a `Name` still needs processing. If so, move it from `remaining` to `pending`. -/ 73 | def isTodo (name : Name) : M Bool := do 74 | let r := (← get).remaining 75 | if r.contains name then 76 | modify fun s => { s with remaining := s.remaining.erase name, pending := s.pending.insert name } 77 | return true 78 | else 79 | return false 80 | 81 | /-- Use the current `Environment` to throw a `Kernel.Exception`. -/ 82 | def throwKernelException (ex : Exception) : M α := do 83 | let ctx := { fileName := "", options := pp.match.set (pp.rawOnError.set {} true) false, fileMap := default } 84 | let state := { env := .ofKernelEnv (← get).env } 85 | Prod.fst <$> (Lean.Core.CoreM.toIO · ctx state) do Lean.throwKernelException ex 86 | 87 | def Lean.Declaration.name : Declaration → String 88 | | .axiomDecl d => s!"axiomDecl {d.name}" 89 | | .defnDecl d => s!"defnDecl {d.name}" 90 | | .thmDecl d => s!"thmDecl {d.name}" 91 | | .opaqueDecl d => s!"opaqueDecl {d.name}" 92 | | .quotDecl => s!"quotDecl" 93 | | .mutualDefnDecl d => s!"mutualDefnDecl {d.map (·.name)}" 94 | | .inductDecl _ _ d _ => s!"inductDecl {d.map (·.name)}" 95 | 96 | /-- Add a declaration, possibly throwing a `KernelException`. -/ 97 | def addDecl (d : Declaration) : M Unit := do 98 | if (← read).verbose then 99 | println! "adding {d.name}" 100 | let t1 ← IO.monoMsNow 101 | match Lean4Lean.addDecl (← get).env d true with 102 | | .ok env => 103 | let t2 ← IO.monoMsNow 104 | if t2 - t1 > 1000 then 105 | if (← read).compare then 106 | let t3 ← match (← get).env.addDecl {} d with 107 | | .ok _ => IO.monoMsNow 108 | | .error ex => _root_.throwKernelException ex 109 | if (t2 - t1) > 2 * (t3 - t2) then 110 | println! 111 | "{(← get).env.header.mainModule}:{d.name}: lean took {t3 - t2}, lean4lean took {t2 - t1}" 112 | else 113 | println! "{(← get).env.header.mainModule}:{d.name}: lean4lean took {t2 - t1}" 114 | else 115 | println! "{(← get).env.header.mainModule}:{d.name}: lean4lean took {t2 - t1}" 116 | modify fun s => { s with env := env } 117 | | .error ex => 118 | throwKernelException ex 119 | 120 | deriving instance BEq for ConstantVal 121 | deriving instance BEq for ConstructorVal 122 | deriving instance BEq for RecursorRule 123 | deriving instance BEq for RecursorVal 124 | 125 | mutual 126 | /-- 127 | Check if a `Name` still needs to be processed (i.e. is in `remaining`). 128 | 129 | If so, recursively replay any constants it refers to, 130 | to ensure we add declarations in the right order. 131 | 132 | The construct the `Declaration` from its stored `ConstantInfo`, 133 | and add it to the environment. 134 | -/ 135 | partial def replayConstant (name : Name) : M Unit := do 136 | if ← isTodo name then 137 | let some ci := (← read).newConstants[name]? | unreachable! 138 | replayConstants ci.getUsedConstants 139 | -- Check that this name is still pending: a mutual block may have taken care of it. 140 | if (← get).pending.contains name then 141 | match ci with 142 | | .defnInfo info => 143 | addDecl (Declaration.defnDecl info) 144 | | .thmInfo info => 145 | addDecl (Declaration.thmDecl info) 146 | | .axiomInfo info => 147 | addDecl (Declaration.axiomDecl info) 148 | | .opaqueInfo info => 149 | addDecl (Declaration.opaqueDecl info) 150 | | .inductInfo info => 151 | let lparams := info.levelParams 152 | let nparams := info.numParams 153 | let all ← info.all.mapM fun n => do pure <| (← read).newConstants[n]! 154 | for o in all do 155 | -- There is exactly one awkward special case here: 156 | -- `String` is a primitive type, which depends on `Char.ofNat` to exist 157 | -- because the kernel treats the existence of the `String` type as license 158 | -- to use string literals, which use `Char.ofNat` internally. However 159 | -- this definition is not transitively reachable from the declaration of `String`. 160 | if o.name == ``String then replayConstant ``Char.ofNat 161 | modify fun s => 162 | { s with remaining := s.remaining.erase o.name, pending := s.pending.erase o.name } 163 | let ctorInfo ← all.mapM fun ci => do 164 | pure (ci, ← ci.inductiveVal!.ctors.mapM fun n => do 165 | pure (← read).newConstants[n]!) 166 | -- Make sure we are really finished with the constructors. 167 | for (_, ctors) in ctorInfo do 168 | for ctor in ctors do 169 | replayConstants ctor.getUsedConstants 170 | let types : List InductiveType := ctorInfo.map fun ⟨ci, ctors⟩ => 171 | { name := ci.name 172 | type := ci.type 173 | ctors := ctors.map fun ci => { name := ci.name, type := ci.type } } 174 | addDecl (Declaration.inductDecl lparams nparams types false) 175 | -- We postpone checking constructors, 176 | -- and at the end make sure they are identical 177 | -- to the constructors generated when we replay the inductives. 178 | | .ctorInfo info => 179 | modify fun s => { s with postponedConstructors := s.postponedConstructors.insert info.name } 180 | -- Similarly we postpone checking recursors. 181 | | .recInfo info => 182 | modify fun s => { s with postponedRecursors := s.postponedRecursors.insert info.name } 183 | | .quotInfo _ => 184 | addDecl (Declaration.quotDecl) 185 | modify fun s => { s with pending := s.pending.erase name } 186 | 187 | /-- Replay a set of constants one at a time. -/ 188 | partial def replayConstants (names : NameSet) : M Unit := do 189 | for n in names do replayConstant n 190 | 191 | end 192 | 193 | /-- 194 | Check that all postponed constructors are identical to those generated 195 | when we replayed the inductives. 196 | -/ 197 | def checkPostponedConstructors : M Unit := do 198 | for ctor in (← get).postponedConstructors do 199 | match (← get).env.constants.find? ctor, (← read).newConstants[ctor]? with 200 | | some (.ctorInfo info), some (.ctorInfo info') => 201 | if ! (info == info') then throw <| IO.userError s!"Invalid constructor {ctor}" 202 | | _, _ => throw <| IO.userError s!"No such constructor {ctor}" 203 | 204 | /-- 205 | Check that all postponed recursors are identical to those generated 206 | when we replayed the inductives. 207 | -/ 208 | def checkPostponedRecursors : M Unit := do 209 | for ctor in (← get).postponedRecursors do 210 | match (← get).env.constants.find? ctor, (← read).newConstants[ctor]? with 211 | | some (.recInfo info), some (.recInfo info') => 212 | if ! (info == info') then throw <| IO.userError s!"Invalid recursor {ctor}" 213 | | _, _ => throw <| IO.userError s!"No such recursor {ctor}" 214 | 215 | /-- 216 | Check that at the end of (any) file, the quotient module is initialized by the end. 217 | (It will already be initialized at the beginning, unless this is the very first file, 218 | `Init.Core`, which is responsible for initializing it.) 219 | This is needed because it is an assumption in `finalizeImport`. 220 | -/ 221 | def checkQuotInit : M Unit := do 222 | unless (← get).env.quotInit do 223 | throw <| IO.userError s!"initial import (Init.Core) didn't initialize quotient module" 224 | 225 | /-- "Replay" some constants into an `Environment`, sending them to the kernel for checking. -/ 226 | def replay (ctx : Context) (env : Environment) (decl : Option Name := none) : IO Environment := do 227 | let mut remaining : NameSet := ∅ 228 | for (n, ci) in ctx.newConstants.toList do 229 | -- We skip unsafe constants, and also partial constants. 230 | -- Later we may want to handle partial constants. 231 | if !ci.isUnsafe && !ci.isPartial then 232 | remaining := remaining.insert n 233 | let (_, s) ← StateRefT'.run (s := { env, remaining }) do 234 | ReaderT.run (r := ctx) do 235 | match decl with 236 | | some d => replayConstant d 237 | | none => 238 | for n in remaining do 239 | replayConstant n 240 | checkPostponedConstructors 241 | checkPostponedRecursors 242 | checkQuotInit 243 | return s.env 244 | 245 | open private ImportedModule.mk from Lean.Environment in 246 | unsafe def replayFromImports (module : Name) (verbose := false) (compare := false) : IO Unit := do 247 | let mFile ← findOLean module 248 | unless (← mFile.pathExists) do 249 | throw <| IO.userError s!"object file '{mFile}' of module {module} does not exist" 250 | let (mod, region) ← readModuleData mFile 251 | let (_, s) ← importModulesCore mod.imports 252 | |>.run (s := { 253 | moduleNameMap := ({} : Std.HashMap ..).insert module 254 | (ImportedModule.mk module (parts := #[(mod, region)])) 255 | moduleNames := #[module] }) 256 | let env ← match Kernel.Environment.finalizeImport s #[{module}] module 0 with 257 | | .ok env => pure env 258 | | .error e => throw <| .userError <| ← (e.toMessageData {}).toString 259 | let mut newConstants := {} 260 | for name in mod.constNames, ci in mod.constants do 261 | newConstants := newConstants.insert name ci 262 | let env' ← replay { newConstants, verbose, compare } env 263 | (Environment.ofKernelEnv env').freeRegions 264 | region.free 265 | 266 | unsafe def replayFromFresh (module : Name) 267 | (verbose := false) (compare := false) (decl : Option Name := none) : IO Unit := do 268 | Lean.withImportModules #[{module}] {} (trustLevel := 0) fun env => do 269 | let ctx := { newConstants := env.constants.map₁, verbose, compare } 270 | discard <| replay ctx (.empty module) decl 271 | 272 | /-- Read the name of the main module from the `lake-manifest`. -/ 273 | -- This has been copied from `ImportGraph.getCurrentModule` in the 274 | -- https://github.com/leanprover-community/import-graph repository. 275 | def getCurrentModule : IO Name := do 276 | match (← Lake.Manifest.load? ⟨"lake-manifest.json"⟩) with 277 | | none => 278 | -- TODO: should this be caught? 279 | pure .anonymous 280 | | some manifest => 281 | -- TODO: This assumes that the `package` and the default `lean_lib` 282 | -- have the same name up to capitalisation. 283 | -- Would be better to read the `.defaultTargets` from the 284 | -- `← getRootPackage` from `Lake`, but I can't make that work with the monads involved. 285 | return manifest.name.capitalize 286 | 287 | /-- 288 | Run as e.g. `lake exe lean4lean` to check everything in the current project. 289 | or e.g. `lake exe lean4lean Mathlib.Data.Nat` to check everything with module name 290 | starting with `Mathlib.Data.Nat`. 291 | 292 | This will replay all the new declarations from the target file into the `Environment` 293 | as it was at the beginning of the file, using the kernel to check them. 294 | 295 | You can also use `lake exe lean4lean --fresh Mathlib.Data.Nat.Basic` to replay all the constants 296 | (both imported and defined in that file) into a fresh environment, 297 | but this can only be used on a single file. 298 | -/ 299 | unsafe def main (args : List String) : IO UInt32 := do 300 | initSearchPath (← findSysroot) 301 | let (flags, args) := args.partition fun s => s.startsWith "-" 302 | let verbose := "-v" ∈ flags || "--verbose" ∈ flags 303 | let fresh : Bool := "--fresh" ∈ flags 304 | let compare : Bool := "--compare" ∈ flags 305 | let targets ← do 306 | match args with 307 | | [] => pure [← getCurrentModule] 308 | | args => args.mapM fun arg => do 309 | let mod := arg.toName 310 | if mod.isAnonymous then 311 | throw <| IO.userError s!"Could not resolve module: {arg}" 312 | else 313 | pure mod 314 | let mut targetModules := #[] 315 | let sp ← searchPathRef.get 316 | for target in targets do 317 | let mut found := false 318 | for path in (← SearchPath.findAllWithExt sp "olean") do 319 | if let some m := (← searchModuleNameOfFileName path sp) then 320 | if target.isPrefixOf m then 321 | targetModules := targetModules.push m 322 | found := true 323 | if not found then 324 | throw <| IO.userError s!"Could not find any oleans for: {target}" 325 | if fresh then 326 | if targetModules.size != 1 then 327 | throw <| IO.userError "--fresh flag is only valid when specifying a single module" 328 | for m in targetModules do 329 | if verbose then IO.println s!"replaying {m} with --fresh" 330 | replayFromFresh m verbose compare 331 | else 332 | let mut tasks := #[] 333 | for m in targetModules do 334 | tasks := tasks.push (m, ← IO.asTask (replayFromImports m verbose compare)) 335 | let mut err := false 336 | for (m, t) in tasks do 337 | if verbose then IO.println s!"replaying {m}" 338 | if let .error e := t.get then 339 | IO.eprintln s!"lean4lean found a problem in {m}:\n{e.toString}" 340 | err := true 341 | if err then return 1 342 | return 0 343 | -------------------------------------------------------------------------------- /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] [--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 | ## (Selected) file breakdown 42 | 43 | * `Main.lean`: command line app 44 | * `Lean4Lean`: source files 45 | * `Environment.lean`: library entry point 46 | * `TypeChecker.lean`: main recursive function 47 | * `Inductive` 48 | * `Add.lean`: constructing inductive recursors 49 | * `Reduce.lean`: inductive iota rules 50 | * `Quot.lean`: quotient types handling 51 | * `UnionFind.lean`: a union-find data structure 52 | * `Std`: stuff that should exist upstream 53 | * `Theory`: lean metatheory 54 | * `VLevel.lean`: level expressions 55 | * `VExpr.lean`: expressions (boring de Bruijn variable theorems are here) 56 | * `VDecl.lean`: declarations 57 | * `VEnv.lean`: environment 58 | * `Meta.lean`: elaborator producing `VExpr`s 59 | * `Inductive.lean`: inductive types 60 | * `Quot.lean`: quotient types 61 | * `Typing` 62 | * `Basic.lean`: The typing relation itself 63 | * `Lemmas.lean`: theorems about the typing relation 64 | * `Meta.lean`: tactic for proving typing judgments 65 | * `Strong.lean`: proof that you can have all the inductive hypotheses 66 | * `Stratified.lean`: (experimental) stratified typing judgment 67 | * `StratifiedUntyped.lean` (experimental) another stratified typing judgment 68 | * `UniqueTyping.lean`: conjectures about the typing relation 69 | * `ParallelReduction.lean`: (experimental) stuff related to church-rosser 70 | * `Env.lean`: typing for environments 71 | * `Verify`: relation between the metatheory and the kernel 72 | * `Axioms.lean`: theorems about upstream opaques that shouldn't be opaque 73 | * `UnionFind.lean`: proof of correctness of union-find 74 | * `VLCtx.lean`: a "translation context" suitable for translating expressions 75 | * `LocalContext.lean`: properties of lean's `LocalContext` type 76 | * `NameGenerator.lean`: properties of the fresh name generator 77 | * `Typing` 78 | * `Expr.lean`: translating expressions (`TrExpr` is here) 79 | * `Lemmas.lean`: properties of `TrExpr` 80 | * `ConditionallyTyped.lean`: properties of expressions in caches that may be out of scope 81 | * `Environment.lean`: translating environments 82 | -------------------------------------------------------------------------------- /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": "7a0d63fbf8fd350e891868a06d9927efa545ac1e", 9 | "name": "batteries", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.20.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.20.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_exe]] 21 | name = "lean4lean" 22 | root = "Main" 23 | supportInterpreter = true 24 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.20.1 2 | --------------------------------------------------------------------------------