Skip to content

Commit fa20b35

Browse files
sonmarchoRyan Lahfa
authored andcommitted
Update the Lean library to 4.11.0-rc2 (#316)
* Make minor updates * Start updating Lean to v4.11.0-rc2 * Fix Diverge/Elab.lean * Make minor modifications * Update the Lean tests and fix the proofs * Add a model for clone::clone::Clone::clone_from to the Lean library * Revert some modifications * Revert more modifications
1 parent ac61031 commit fa20b35

File tree

13 files changed

+139
-536
lines changed

13 files changed

+139
-536
lines changed

backends/lean/Base/Diverge/Elab.lean

Lines changed: 90 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -10,9 +10,6 @@ namespace Diverge
1010
/- Automating the generation of the encoding and the proofs so as to use nice
1111
syntactic sugar. -/
1212

13-
syntax (name := divergentDef)
14-
declModifiers "divergent" "def" declId ppIndent(optDeclSig) declVal : command
15-
1613
open Lean Elab Term Meta Primitives Lean.Meta
1714
open Utils
1815

@@ -1389,17 +1386,43 @@ def addPreDefinitions (preDefs : Array PreDefinition) : TermElabM Unit := withLC
13891386
else return ()
13901387
catch _ => s.restore
13911388

1392-
-- The following two functions are copy-pasted from Lean.Elab.MutualDef
1393-
1389+
-- The following three functions are copy-pasted from Lean.Elab.MutualDef.lean
13941390
open private elabHeaders levelMVarToParamHeaders getAllUserLevelNames withFunLocalDecls elabFunValues
13951391
instantiateMVarsAtHeader instantiateMVarsAtLetRecToLift checkLetRecsToLiftTypes withUsed from Lean.Elab.MutualDef
13961392

1393+
-- Copy/pasted from Lean.Elab.Term.withHeaderSecVars (because the definition is private)
1394+
private def Term.withHeaderSecVars {α} (vars : Array Expr) (includedVars : List Name) (headers : Array DefViewElabHeader)
1395+
(k : Array Expr → TermElabM α) : TermElabM α := do
1396+
let (_, used) ← collectUsed.run {}
1397+
let (lctx, localInsts, vars) ← removeUnused vars used
1398+
withLCtx lctx localInsts <| k vars
1399+
where
1400+
collectUsed : StateRefT CollectFVars.State MetaM Unit := do
1401+
-- directly referenced in headers
1402+
headers.forM (·.type.collectFVars)
1403+
-- included by `include`
1404+
vars.forM fun var => do
1405+
let ldecl ← getFVarLocalDecl var
1406+
if includedVars.contains ldecl.userName then
1407+
modify (·.add ldecl.fvarId)
1408+
-- transitively referenced
1409+
get >>= (·.addDependencies) >>= set
1410+
-- instances (`addDependencies` unnecessary as by definition they may only reference variables
1411+
-- already included)
1412+
vars.forM fun var => do
1413+
let ldecl ← getFVarLocalDecl var
1414+
let st ← get
1415+
if ldecl.binderInfo.isInstImplicit && (← getFVars ldecl.type).all st.fvarSet.contains then
1416+
modify (·.add ldecl.fvarId)
1417+
getFVars (e : Expr) : MetaM (Array FVarId) :=
1418+
(·.2.fvarIds) <$> e.collectFVars.run {}
1419+
13971420
-- Comes from Term.isExample
13981421
def isExample (views : Array DefView) : Bool :=
13991422
views.any (·.kind.isExample)
14001423

14011424
open Language in
1402-
def Term.elabMutualDef (vars : Array Expr) (views : Array DefView) : TermElabM Unit :=
1425+
def Term.elabMutualDef (vars : Array Expr) (includedVars : List Name) (views : Array DefView) : TermElabM Unit :=
14031426
if isExample views then
14041427
withoutModifyingEnv do
14051428
-- save correct environment in info tree
@@ -1418,7 +1441,7 @@ where
14181441
withFunLocalDecls headers fun funFVars => do
14191442
for view in views, funFVar in funFVars do
14201443
addLocalVarInfo view.declId funFVar
1421-
-- Modification 1:
1444+
-- MODIFICATION 1:
14221445
-- Add fake use site to prevent "unused variable" warning (if the
14231446
-- function is actually not recursive, Lean would print this warning).
14241447
-- Remark: we could detect this case and encode the function without
@@ -1428,7 +1451,7 @@ where
14281451
addTermInfo' view.declId funFVar
14291452
let values ←
14301453
try
1431-
let values ← elabFunValues headers
1454+
let values ← elabFunValues headers vars includedVars
14321455
Term.synthesizeSyntheticMVarsNoPostponing
14331456
values.mapM (instantiateMVars ·)
14341457
catch ex =>
@@ -1438,18 +1461,23 @@ where
14381461
let letRecsToLift ← getLetRecsToLift
14391462
let letRecsToLift ← letRecsToLift.mapM instantiateMVarsAtLetRecToLift
14401463
checkLetRecsToLiftTypes funFVars letRecsToLift
1441-
withUsed vars headers values letRecsToLift fun vars => do
1464+
(if headers.all (·.kind.isTheorem) && !deprecated.oldSectionVars.get (← getOptions) then withHeaderSecVars vars includedVars headers else withUsed vars headers values letRecsToLift) fun vars => do
14421465
let preDefs ← MutualClosure.main vars headers funFVars values letRecsToLift
14431466
for preDef in preDefs do
14441467
trace[Elab.definition] "{preDef.declName} : {preDef.type} :=\n{preDef.value}"
14451468
let preDefs ← withLevelNames allUserLevelNames <| levelMVarToParamPreDecls preDefs
14461469
let preDefs ← instantiateMVarsAtPreDecls preDefs
1470+
let preDefs ← shareCommonPreDefs preDefs
14471471
let preDefs ← fixLevelParams preDefs scopeLevelNames allUserLevelNames
14481472
for preDef in preDefs do
14491473
trace[Elab.definition] "after eraseAuxDiscr, {preDef.declName} : {preDef.type} :=\n{preDef.value}"
14501474
checkForHiddenUnivLevels allUserLevelNames preDefs
1451-
addPreDefinitions preDefs -- Modification 2: we use our custom function here
1475+
addPreDefinitions preDefs -- MODIFICATION 2: we use our custom function here
14521476
processDeriving headers
1477+
for view in views, header in headers do
1478+
-- NOTE: this should be the full `ref`, and thus needs to be done after any snapshotting
1479+
-- that depends only on a part of the ref
1480+
addDeclarationRanges header.declName view.ref
14531481

14541482
processDeriving (headers : Array DefViewElabHeader) := do
14551483
for header in headers, view in views do
@@ -1460,22 +1488,61 @@ where
14601488
unless (← processDefDeriving className header.declName) do
14611489
throwError "failed to synthesize instance '{className}' for '{header.declName}'"
14621490

1491+
#check Command.elabMutualDef
1492+
1493+
-- Copy/pasted from Lean.Elab.MutualDef
14631494
open Command in
1495+
open Language in
14641496
def Command.elabMutualDef (ds : Array Syntax) : CommandElabM Unit := do
1465-
let views ← ds.mapM fun d => do
1466-
let `($mods:declModifiers divergent def $id:declId $sig:optDeclSig $val:declVal) := d
1467-
| throwUnsupportedSyntax
1468-
let modifiers ← elabModifiers mods
1469-
let (binders, type) := expandOptDeclSig sig
1470-
let deriving? := none
1471-
let headerRef := Syntax.missing -- Not sure what to put here
1472-
pure { ref := d, kind := DefKind.def, headerRef, modifiers,
1473-
declId := id, binders, type? := type, value := val, deriving? }
1474-
runTermElabM fun vars => Term.elabMutualDef vars views
1497+
let opts ← getOptions
1498+
withAlwaysResolvedPromises ds.size fun headerPromises => do
1499+
let snap? := (← read).snap?
1500+
let mut views := #[]
1501+
let mut defs := #[]
1502+
let mut reusedAllHeaders := true
1503+
for h : i in [0:ds.size], headerPromise in headerPromises do
1504+
let d := ds[i]
1505+
let modifiers ← elabModifiers d[0]
1506+
if ds.size > 1 && modifiers.isNonrec then
1507+
throwErrorAt d "invalid use of 'nonrec' modifier in 'mutual' block"
1508+
let mut view ← mkDefView modifiers d[2] -- MODIFICATION: changed the index to 2
1509+
let fullHeaderRef := mkNullNode #[d[0], view.headerRef]
1510+
if let some snap := snap? then
1511+
view := { view with headerSnap? := some {
1512+
old? := do
1513+
-- transitioning from `Context.snap?` to `DefView.headerSnap?` invariant: if the
1514+
-- elaboration context and state are unchanged, and the syntax of this as well as all
1515+
-- previous headers is unchanged, then the elaboration result for this header (which
1516+
-- includes state from elaboration of previous headers!) should be unchanged.
1517+
guard reusedAllHeaders
1518+
let old ← snap.old?
1519+
-- blocking wait, `HeadersParsedSnapshot` (and hopefully others) should be quick
1520+
let old ← old.val.get.toTyped? DefsParsedSnapshot
1521+
let oldParsed ← old.defs[i]?
1522+
guard <| fullHeaderRef.eqWithInfoAndTraceReuse opts oldParsed.fullHeaderRef
1523+
-- no syntax guard to store, we already did the necessary checks
1524+
return ⟨.missing, oldParsed.headerProcessedSnap⟩
1525+
new := headerPromise
1526+
} }
1527+
defs := defs.push {
1528+
fullHeaderRef
1529+
headerProcessedSnap := { range? := d.getRange?, task := headerPromise.result }
1530+
}
1531+
reusedAllHeaders := reusedAllHeaders && view.headerSnap?.any (·.old?.isSome)
1532+
views := views.push view
1533+
if let some snap := snap? then
1534+
-- no non-fatal diagnostics at this point
1535+
snap.new.resolve <| .ofTyped { defs, diagnostics := .empty : DefsParsedSnapshot }
1536+
let includedVars := (← getScope).includedVars
1537+
runTermElabM fun vars => Term.elabMutualDef vars includedVars views
1538+
1539+
syntax (name := divergentDef)
1540+
declModifiers "divergent" Lean.Parser.Command.definition : command
14751541

14761542
-- Special command so that we don't fall back to the built-in mutual when we produce an error.
14771543
local syntax "_divergent" Parser.Command.mutual : command
1478-
elab_rules : command | `(_divergent mutual $decls* end) => Command.elabMutualDef decls
1544+
elab_rules : command
1545+
| `(_divergent mutual $decls* end) => Command.elabMutualDef decls
14791546

14801547
macro_rules
14811548
| `(mutual $decls* end) => do
@@ -1501,6 +1568,8 @@ namespace Tests
15011568

15021569
/- Some examples of partial functions -/
15031570

1571+
-- set_option trace.Diverge true
1572+
-- set_option pp.rawOnError true
15041573
--set_option trace.Diverge.def true
15051574
--set_option trace.Diverge.def.genBody true
15061575
--set_option trace.Diverge.def.valid true

backends/lean/Base/Primitives/Core.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -74,3 +74,10 @@ end core
7474

7575
/- [core::option::Option::is_none] -/
7676
@[simp] def core.option.Option.is_none (T: Type) (self: Option T): Bool := self.isNone
77+
78+
/- [core::clone::Clone::clone_from]:
79+
Source: '/rustc/library/core/src/clone.rs', lines 175:4-175:43
80+
Name pattern: core::clone::Clone::clone_from -/
81+
@[simp] def core.clone.Clone.clone_from
82+
{Self : Type} (cloneInst : core.clone.Clone Self) (_self : Self) (source : Self) : Result Self :=
83+
cloneInst.clone source

backends/lean/Base/Utils.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -869,7 +869,7 @@ def evalAesopSaturate (options : Aesop.Options') (ruleSets : Array Name) : Tacti
869869
let rss ← Aesop.Frontend.getGlobalRuleSets ruleSets
870870
let rs ← Aesop.mkLocalRuleSet rss options
871871
|> Aesop.ElabM.runForwardElab (← getMainGoal)
872-
tryLiftMetaTactic1 (Aesop.saturate rs · |>.run { options }) "Aesop.saturate failed"
872+
tryLiftMetaTactic1 (Aesop.saturate rs · options) "Aesop.saturate failed"
873873

874874
/-- Normalize the let-bindings by inlining them -/
875875
def normalizeLetBindings (e : Expr) : MetaM Expr :=

backends/lean/lake-manifest.json

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@
55
"type": "git",
66
"subDir": null,
77
"scope": "leanprover-community",
8-
"rev": "c0efc1fd2a0bec51bd55c5b17348af13d7419239",
8+
"rev": "e6d3a32d66252a70fda1d56463e1da975b3b8f53",
99
"name": "batteries",
1010
"manifestFile": "lake-manifest.json",
1111
"inputRev": "main",
@@ -15,7 +15,7 @@
1515
"type": "git",
1616
"subDir": null,
1717
"scope": "leanprover-community",
18-
"rev": "a7bfa63f5dddbcab2d4e0569c4cac74b2585e2c6",
18+
"rev": "71f54425e6fe0fa75f3aef33a2813a7898392222",
1919
"name": "Qq",
2020
"manifestFile": "lake-manifest.json",
2121
"inputRev": "master",
@@ -25,7 +25,7 @@
2525
"type": "git",
2626
"subDir": null,
2727
"scope": "leanprover-community",
28-
"rev": "cf30d04b6448dbb5a5b30a7d031e3949e74b9dd1",
28+
"rev": "c792cfd1efe6e01cb176e158ddb195bedfb7ad33",
2929
"name": "aesop",
3030
"manifestFile": "lake-manifest.json",
3131
"inputRev": "master",
@@ -35,27 +35,27 @@
3535
"type": "git",
3636
"subDir": null,
3737
"scope": "leanprover-community",
38-
"rev": "d1b33202c3a29a079f292de65ea438648123b635",
38+
"rev": "a96aee5245720f588876021b6a0aa73efee49c76",
3939
"name": "proofwidgets",
4040
"manifestFile": "lake-manifest.json",
41-
"inputRev": "v0.0.39",
41+
"inputRev": "v0.0.41",
4242
"inherited": true,
4343
"configFile": "lakefile.lean"},
4444
{"url": "https://github.com/leanprover/lean4-cli",
4545
"type": "git",
4646
"subDir": null,
4747
"scope": "",
48-
"rev": "a11566029bd9ec4f68a65394e8c3ff1af74c1a29",
48+
"rev": "2cf1030dc2ae6b3632c84a09350b675ef3e347d0",
4949
"name": "Cli",
5050
"manifestFile": "lake-manifest.json",
5151
"inputRev": "main",
5252
"inherited": true,
53-
"configFile": "lakefile.lean"},
53+
"configFile": "lakefile.toml"},
5454
{"url": "https://github.com/leanprover-community/import-graph",
5555
"type": "git",
5656
"subDir": null,
5757
"scope": "leanprover-community",
58-
"rev": "68b518c9b352fbee16e6d632adcb7a6d0760e2b7",
58+
"rev": "57bd2065f1dbea5e9235646fb836c7cea9ab03b6",
5959
"name": "importGraph",
6060
"manifestFile": "lake-manifest.json",
6161
"inputRev": "main",
@@ -65,7 +65,7 @@
6565
"type": "git",
6666
"subDir": null,
6767
"scope": "",
68-
"rev": "e242f1edcacf917f40fae9b81f57f4bd0a4e45ac",
68+
"rev": "9d7806d77c33a5e19050de8fbdc106a28150ec71",
6969
"name": "mathlib",
7070
"manifestFile": "lake-manifest.json",
7171
"inputRev": null,

backends/lean/lean-toolchain

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1 +1 @@
1-
leanprover/lean4:v4.10.0-rc2
1+
leanprover/lean4:v4.11.0-rc2

compiler/ExtractBuiltin.ml

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -508,6 +508,7 @@ let mk_builtin_funs () : (pattern * bool list option * builtin_fun_info) list =
508508
~extract_name:
509509
(Some (backend_choice "" "core::option::Option::is_none"))
510510
~can_fail:false ();
511+
mk_fun "core::clone::Clone::clone_from" ();
511512
]
512513
@ List.map
513514
(fun ty ->

tests/lean/Avl/Properties.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -799,7 +799,6 @@ theorem Tree.insert_in_opt_node_spec
799799
split_conjs
800800
. simp [Node.invAux, Node.balanceFactor]
801801
. simp [Subtree.inv]
802-
. apply Set.ext; simp
803802
. -- tree = some
804803
rename Node T => node
805804
have hNodeInv : Node.inv node := by simp_all (config := {maxDischargeDepth := 1})

tests/lean/External/FunsExternal_Template.lean

Lines changed: 0 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -24,10 +24,3 @@ axiom core.cell.Cell.get_mut
2424
core.cell.Cell T → State → Result (State × (T × (T → State → Result
2525
(State × (core.cell.Cell T)))))
2626

27-
/- [core::clone::Clone::clone_from]:
28-
Source: '/rustc/library/core/src/clone.rs', lines 175:4-175:43
29-
Name pattern: core::clone::Clone::clone_from -/
30-
axiom core.clone.Clone.clone_from
31-
{Self : Type} (self_clause : core.clone.Clone Self) :
32-
Self → Self → State → Result (State × Self)
33-

tests/lean/Tutorial/Exercises.lean

Lines changed: 6 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -165,7 +165,7 @@ theorem even_spec (n : U32) :
165165
. progress as ⟨ n' ⟩
166166
progress as ⟨ b ⟩
167167
simp [*]
168-
simp [Int.even_sub]
168+
simp [Int.odd_sub]
169169
termination_by n.toNat
170170
decreasing_by scalar_decr_tac
171171

@@ -437,10 +437,9 @@ divergent def add_no_overflow_loop
437437
- `scalar_eq_nf`: similar, but tuned to prove goals of the shape: `... = ...`
438438
You can try both tactics and see their effect.
439439
-/
440-
-- Here, we're using ring_nf
441440
@[simp]
442-
theorem toInt_aux_append (l0 l1 : List U32) :
443-
toInt_aux (l0 ++ l1) = toInt_aux l0 + 2 ^ (32 * l0.length) * toInt_aux l1 := by
441+
theorem toInt_aux_drop (l : List U32) (i : Nat) (h0 : i < l.length) :
442+
toInt_aux (l.drop i) = l.index i + 2 ^ 32 * toInt_aux (l.drop (i + 1)) := by
444443
sorry
445444

446445
/-- You will need this lemma for the proof of `add_no_overflow_loop_spec`.
@@ -461,16 +460,11 @@ theorem toInt_aux_update (l : List U32) (i : Nat) (x : U32) (h0 : i < l.length)
461460
toInt_aux (l.update i x) = toInt_aux l + 2 ^ (32 * i) * (x - l.index i) := by
462461
sorry
463462

464-
/-- You will need this lemma for the proof of `add_no_overflow_loop_spec`.
463+
/-- The proof about `add_no_overflow_loop`.
465464
466-
Advice: do the proof of `add_no_overflow_loop_spec` first, then come back to prove this lemma.
465+
Hint: you will need to reason about non-linear arithmetic with `scalar_nf` and
466+
`scalar_eq_nf`` (see above).
467467
-/
468-
@[simp]
469-
theorem toInt_aux_drop (l : List U32) (i : Nat) (h0 : i < l.length) :
470-
toInt_aux (l.drop i) = l.index i + 2 ^ 32 * toInt_aux (l.drop (i + 1)) := by
471-
sorry
472-
473-
/-- The proof about `add_no_overflow_loop` -/
474468
@[pspec]
475469
theorem add_no_overflow_loop_spec
476470
(x : alloc.vec.Vec U32) (y : alloc.vec.Vec U32) (i : Usize)

0 commit comments

Comments
 (0)