From 7823d59b2d9436f1bf0042fff62ee2bcc4c29ec0 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 3 Jul 2018 18:57:57 +0200 Subject: Refactor HoTT_Methods.thy, proved more stuff with new methods. --- HoTT_Methods.thy | 92 +++++++++----------------------------------------------- 1 file changed, 14 insertions(+), 78 deletions(-) (limited to 'HoTT_Methods.thy') diff --git a/HoTT_Methods.thy b/HoTT_Methods.thy index bce5123..20f3ece 100644 --- a/HoTT_Methods.thy +++ b/HoTT_Methods.thy @@ -2,8 +2,7 @@ Author: Josh Chen Date: Jun 2018 -Method setup for the HoTT library. -Relies heavily on Eisbach. +Method setup for the HoTT library. Relies heavily on Eisbach. *) theory HoTT_Methods @@ -11,14 +10,14 @@ theory HoTT_Methods "HOL-Eisbach.Eisbach" "HOL-Eisbach.Eisbach_Tools" HoTT_Base - Equal - Prod - Sum begin + section \Method setup\ -text "Prove type judgments \A : U\ and inhabitation judgments \a : A\ using rules declared [intro] and [elim], as well as additional provided lemmas." +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)+ @@ -33,83 +32,20 @@ method wellformed uses jdgmt = ( ) -text "Combine \simple\ and \wellformed\ to search for derivations." +text "Combine \simple\ and \wellformed\ to search for derivations. +\wellformed\ uses the facts passed as \lems\ to derive any required typing judgments. +Definitions passed as \unfolds\ are unfolded throughout. + +Roughly speaking \derive\ is more powerful than \simple\, but may fail to find a proof where \simple\ does if it reduces too eagerly." -method derive uses lems = ( - catch \unfold lems\ \fail\ | - catch \simple lems: lems\ \fail\ | +method derive uses lems unfolds = ( + unfold unfolds | + simple lems: lems | match lems in lem: "?X : ?Y" \ \wellformed jdgmt: lem\ )+ -section \Examples\ - -lemma - assumes "A : U" "B: A \ U" "\x. x : A \ C x: B x \ U" - shows "\x:A. \y:B x. \z:C x y. \w:A. x =\<^sub>A w : U" -by (simple lems: assms) - -lemma - assumes "f : \x:A. \y: B x. \z: C x y. D x y z" - shows - "A : U" and - "B: A \ U" and - "\x. x : A \ C x: B x \ U" and - "\x y. \x : A; y : B x\ \ D x y: C x y \ U" -proof - - show "A : U" by (wellformed jdgmt: assms) - show "B: A \ U" by (wellformed jdgmt: assms) - show "\x. x : A \ C x: B x \ U" by (wellformed jdgmt: assms) - show "\x y. \x : A; y : B x\ \ D x y: C x y \ U" by (wellformed jdgmt: assms) -qed - - -section \Experimental methods\ - -text "Playing around with ML, following CTT/CTT.thy by Larry Paulson." - -lemmas forms = Prod_form Sum_form Equal_form -lemmas intros = Prod_intro Sum_intro Equal_intro -lemmas elims = Prod_elim Sum_elim Equal_elim -lemmas elements = intros elims - -ML \ -(* Try solving \a : A\ by assumption provided \a\ is rigid *) -fun test_assume_tac ctxt = let - fun is_rigid (Const(@{const_name is_of_type},_) $ a $ _) = not(is_Var (head_of a)) - | is_rigid (Const(@{const_name is_a_type},_) $ a $ _ $ _) = not(is_Var (head_of a)) - | is_rigid _ = false - in - SUBGOAL (fn (prem, i) => - if is_rigid (Logic.strip_assums_concl prem) - then assume_tac ctxt i else no_tac) - end - -fun ASSUME ctxt tf i = test_assume_tac ctxt i ORELSE tf i - -(* Solve all subgoals \A : U\ using formation rules. *) -val form_net = Tactic.build_net @{thms forms}; -fun form_tac ctxt = - REPEAT_FIRST (ASSUME ctxt (filt_resolve_from_net_tac ctxt 1 form_net)); - -(* Try to prove inhabitation judgments \a : A\ (\a\ flexible, \A\ rigid) by introduction rules. *) -fun intro_tac ctxt thms = - let val tac = - filt_resolve_from_net_tac ctxt 1 - (Tactic.build_net (thms @ @{thms forms} @ @{thms intros})) - in REPEAT_FIRST (ASSUME ctxt tac) end - -(*Type checking: solve \a : A\ (\a\ rigid, \A\ flexible) by formation, introduction and elimination rules. *) -fun typecheck_tac ctxt thms = - let val tac = - filt_resolve_from_net_tac ctxt 3 - (Tactic.build_net (thms @ @{thms forms} @ @{thms elements})) - in REPEAT_FIRST (ASSUME ctxt tac) end -\ - -method_setup form = \Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\ -method_setup intro = \Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (intro_tac ctxt ths))\ -method_setup typecheck = \Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (typecheck_tac ctxt ths))\ +text "Simplify function applications." end \ No newline at end of file -- cgit v1.2.3