(* Title: HoTT/HoTT_Methods.thy Author: Josh Chen Date: Jun 2018 Method setup for the HoTT library. Relies heavily on Eisbach. *) theory HoTT_Methods imports "HOL-Eisbach.Eisbach" "HOL-Eisbach.Eisbach_Tools" HoTT_Base begin section \Setup\ text "Import the \subst\ method, used by \simplify\." ML_file "~~/src/Tools/misc_legacy.ML" ML_file "~~/src/Tools/IsaPlanner/isand.ML" ML_file "~~/src/Tools/IsaPlanner/rw_inst.ML" ML_file "~~/src/Tools/IsaPlanner/zipper.ML" ML_file "~~/src/Tools/eqsubst.ML" section \Method definitions\ subsection \Simple type rule proof search\ text "Prove type judgments \A : U\ and inhabitation judgments \a : A\ using rules declared [intro] and [elim], as well as additional provided lemmas. Can also perform typechecking, and search for elements of a type." method simple uses lems = (assumption | standard | rule lems)+ subsection \Wellformedness checker\ text "Find a proof of any valid typing judgment derivable from a given wellformed judgment. The named theorem \wellform\ is declared in HoTT_Base.thy." method wellformed uses jdgmt declares wellform = match wellform in rl: "PROP ?P" \ \( catch \rule rl[OF jdgmt]\ \fail\ | catch \wellformed jdgmt: rl[OF jdgmt]\ \fail\ )\ subsection \Derivation search\ text "Combine \simple\ and \wellformed\ to search for derivations of judgments. \wellformed\ uses the facts passed as \lems\ to derive any required typing judgments. Definitions passed as \unfolds\ are unfolded throughout." method derive uses lems unfolds = ( unfold unfolds | simple lems: lems | match lems in lem: "?X : ?Y" \ \wellformed jdgmt: lem\ )+ subsection \Simplification\ text "The methods \rsimplify\ and \simplify\ attempt to solve definitional equations by simplifying lambda applications. \simplify\ is allowed to use the Pure Simplifier and is hence unsuitable for simplifying lambda expressions using only the type rules. In this case use \rsimplify\ instead. Since these methods use \derive\, it is often (but not always) the case that theorems provable with \derive\ are also provable with them. (Whether this is the case depends on whether the call to \subst (0) comp\ fails.)" method rsimplify uses lems = (subst (0) comp, derive lems: lems)+ \ \\subst\ performs an equational rewrite according to some computation rule declared as [comp], after which \derive\ attempts to solve any new assumptions that arise.\ method simplify uses lems = (simp add: lems | rsimplify lems: lems)+ end