(* Title: HoTT/HoTT_Methods.thy Author: Josh Chen 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 \Deriving typing judgments\ method routine uses lems = (assumption | rule lems | standard)+ text " @{method routine} proves routine type judgments \a : A\ using the rules declared [intro] in the respective theory files, as well as additional provided lemmas. " text " \wellformed'\ finds a proof of any valid typing judgment derivable from the judgment passed as \jdmt\. If no judgment is passed, it will try to resolve with the theorems declared \wellform\. \wellform\ is like \wellformed'\ but takes multiple judgments. The named theorem \wellform\ is declared in HoTT_Base.thy. " method wellformed' uses jdmt declares wellform = match wellform in rl: "PROP ?P" \ \( catch \rule rl[OF jdmt]\ \fail\ | catch \wellformed' jdmt: rl[OF jdmt]\ \fail\ )\ method wellformed uses lems = match lems in lem: "?X : ?Y" \ \wellformed' jdmt: lem\ section \Substitution and simplification\ text "Import the \subst\ method, used for substituting definitional equalities." 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" text "Perform basic single-step computations:" method compute uses lems = (subst lems comp | rule lems comp) section \Derivation search\ text " Combine \routine\, \wellformed\, and \compute\ to search for derivations of judgments." method derive uses lems = (routine lems: lems | compute lems: lems | wellformed lems: lems)+ section \Induction\ text " Placeholder section for the automation of induction, i.e. using the elimination rules. At the moment one must directly apply them with \rule X_elim\. " end