|
|
|
@ -264,32 +264,50 @@ elab "MakeGame" : command => do
|
|
|
|
|
if game.worlds.hasLoops then
|
|
|
|
|
throwError "World graph has loops!"
|
|
|
|
|
|
|
|
|
|
-- Compute which tactics are available in which level:
|
|
|
|
|
-- Compute which tactics/lemmas are available in which level:
|
|
|
|
|
let mut newTacticsInWorld : HashMap Name (HashSet Name) := {}
|
|
|
|
|
let mut allTactics : HashSet Name := {}
|
|
|
|
|
let mut newLemmasInWorld : HashMap Name (HashSet Name) := {}
|
|
|
|
|
let mut allLemmas : HashSet Name := {}
|
|
|
|
|
for (worldId, world) in game.worlds.nodes.toArray do
|
|
|
|
|
let mut newTactics : HashSet Name:= {}
|
|
|
|
|
let mut newLemmas : HashSet Name:= {}
|
|
|
|
|
for (_, level) in world.levels.toArray do
|
|
|
|
|
newTactics := newTactics.insertMany level.newTactics
|
|
|
|
|
allTactics := allTactics.insertMany level.newTactics
|
|
|
|
|
newLemmas := newLemmas.insertMany level.newLemmas
|
|
|
|
|
allLemmas := allLemmas.insertMany level.newLemmas
|
|
|
|
|
newTacticsInWorld := newTacticsInWorld.insert worldId newTactics
|
|
|
|
|
newLemmasInWorld := newLemmasInWorld.insert worldId newLemmas
|
|
|
|
|
|
|
|
|
|
let tacticAvailability₀ : HashMap Name TacticAvailability :=
|
|
|
|
|
-- Basic tactic/lemma availability: all locked, none disabled.
|
|
|
|
|
let Availability₀ : HashMap Name Availability :=
|
|
|
|
|
HashMap.ofList $
|
|
|
|
|
allTactics.toList.map fun name =>
|
|
|
|
|
(name, {name, locked := true, disabled := false})
|
|
|
|
|
let lemmaAvailability₀ : HashMap Name Availability :=
|
|
|
|
|
HashMap.ofList $
|
|
|
|
|
allLemmas.toList.map fun name =>
|
|
|
|
|
(name, {name, locked := true, disabled := false})
|
|
|
|
|
|
|
|
|
|
let mut tacticsInWorld : HashMap Name (HashMap Name TacticAvailability) := {}
|
|
|
|
|
-- Availability after a given world
|
|
|
|
|
let mut tacticsInWorld : HashMap Name (HashMap Name Availability) := {}
|
|
|
|
|
let mut lemmasInWorld : HashMap Name (HashMap Name Availability) := {}
|
|
|
|
|
for (worldId, _) in game.worlds.nodes.toArray do
|
|
|
|
|
let mut tactics : HashMap Name TacticAvailability := tacticAvailability₀
|
|
|
|
|
let mut tactics := Availability₀
|
|
|
|
|
let mut lemmas := lemmaAvailability₀
|
|
|
|
|
let predecessors := game.worlds.predecessors worldId
|
|
|
|
|
for predWorldId in predecessors do
|
|
|
|
|
for tac in newTacticsInWorld.find! predWorldId do
|
|
|
|
|
tactics := tactics.insert tac {name := tac, locked := false, disabled := false}
|
|
|
|
|
for lem in newLemmasInWorld.find! predWorldId do
|
|
|
|
|
lemmas := lemmas.insert lem {name := lem, locked := false, disabled := false}
|
|
|
|
|
tacticsInWorld := tacticsInWorld.insert worldId tactics
|
|
|
|
|
lemmasInWorld := lemmasInWorld.insert worldId lemmas
|
|
|
|
|
|
|
|
|
|
for (worldId, world) in game.worlds.nodes.toArray do
|
|
|
|
|
let mut tactics := tacticsInWorld.find! worldId
|
|
|
|
|
let mut lemmas := lemmasInWorld.find! worldId
|
|
|
|
|
|
|
|
|
|
let levels := world.levels.toArray.insertionSort fun a b => a.1 < b.1
|
|
|
|
|
|
|
|
|
@ -298,9 +316,20 @@ elab "MakeGame" : command => do
|
|
|
|
|
tactics := tactics.insert tac {name := tac, locked := false, disabled := false}
|
|
|
|
|
for tac in level.disabledTactics do
|
|
|
|
|
tactics := tactics.insert tac {name := tac, locked := false, disabled := true}
|
|
|
|
|
for lem in level.newLemmas do
|
|
|
|
|
lemmas := lemmas.insert lem {name := lem, locked := false, disabled := false}
|
|
|
|
|
for lem in level.disabledLemmas do
|
|
|
|
|
lemmas := lemmas.insert lem {name := lem, locked := false, disabled := true}
|
|
|
|
|
|
|
|
|
|
let tacticArray := tactics.toArray
|
|
|
|
|
|>.insertionSort (fun a b => a.1.toString < b.1.toString)
|
|
|
|
|
|>.map (·.2)
|
|
|
|
|
|
|
|
|
|
let lemmaArray := lemmas.toArray
|
|
|
|
|
|>.insertionSort (fun a b => a.1.toString < b.1.toString)
|
|
|
|
|
|>.map (·.2)
|
|
|
|
|
|
|
|
|
|
modifyLevel ⟨← getCurGameId, worldId, levelId⟩ fun level => do
|
|
|
|
|
return {level with tactics := tacticArray}
|
|
|
|
|
return {level with
|
|
|
|
|
tactics := tacticArray
|
|
|
|
|
lemmas := lemmaArray}
|
|
|
|
|