From 8f47d9f3ca38a3fc2c3c462d1157866081102ce1 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 25 May 2020 15:24:20 +0200 Subject: 1. equality method now uses general elimination tactic. 2. New constant `` refers to rhs of equalities. --- spartan/theories/Equivalence.thy | 6 +- spartan/theories/Identity.thy | 140 +++++++++++++++++++++++++-------------- 2 files changed, 92 insertions(+), 54 deletions(-) diff --git a/spartan/theories/Equivalence.thy b/spartan/theories/Equivalence.thy index 73f1e0d..3415249 100644 --- a/spartan/theories/Equivalence.thy +++ b/spartan/theories/Equivalence.thy @@ -86,7 +86,7 @@ Lemma (derive) commute_homotopy: shows "(H x) \ g[p] = f[p] \ (H y)" \ \Need this assumption unfolded for typechecking:\ supply assms(8)[unfolded homotopy_def] - apply (equality \p:_\) + apply (eq p) focus vars x apply reduce \ \Here it would really be nice to have automation for transport and @@ -366,7 +366,7 @@ Lemma assumes "A: U i" "B: U i" "p: A =\<^bsub>U i\<^esub> B" shows "A \ B" - by (equality \p:_\) (rule equivalence_refl) + by (eq p) (rule equivalence_refl) text \ The following proof is wordy because (1) the typechecker doesn't rewrite, and @@ -383,7 +383,7 @@ Lemma (derive) id_imp_equiv: \ \Switch off auto-typechecking, which messes with universe levels\ supply [[auto_typechk=false]] - \ apply (equality \p:_\) + \ apply (eq p, typechk) \<^item> prems vars A B apply (rewrite at A in "A \ B" id_comp[symmetric]) apply (rewrite at B in "_ \ B" id_comp[symmetric]) diff --git a/spartan/theories/Identity.thy b/spartan/theories/Identity.thy index 3f16f4c..1ab12ff 100644 --- a/spartan/theories/Identity.thy +++ b/spartan/theories/Identity.thy @@ -40,16 +40,33 @@ axiomatization where lemmas [intros] = IdF IdI and - [elims "?p" (*"?x" "?y"*)] = IdE and + [elims "?p" "?a" "?b"] = IdE and [comps] = Id_comp -section \Induction\ - -ML_file \../lib/equality.ML\ - -method_setup equality = \Scan.lift Parse.thm >> (fn (fact, _) => fn ctxt => - CONTEXT_METHOD (K (Equality.equality_context_tac fact ctxt)))\ +section \Path induction\ + +method_setup eq = \ +Args.term >> (fn tm => fn ctxt => CONTEXT_METHOD (K ( + CONTEXT_SUBGOAL (fn (goal, i) => + let + val facts = Proof_Context.facts_of ctxt + val prems = Logic.strip_assums_hyp goal + val template = \<^const>\has_type\ + $ tm + $ (\<^const>\Id\ $ Var (("*?A", 0), \<^typ>\o\) + $ Var (("*?x", 0), \<^typ>\o\) + $ Var (("*?y", 0), \<^typ>\o\)) + val types = + map (Thm.prop_of o #1) (Facts.could_unify facts template) + @ filter (fn prem => Term.could_unify (template, prem)) prems + |> map Lib.type_of_typing + in case types of + (\<^const>\Id\ $ _ $ x $ y)::_ => + elim_context_tac [tm, x, y] ctxt i + | _ => Context_Tactic.CONTEXT_TACTIC no_tac + end) 1))) +\ section \Symmetry and transitivity\ @@ -57,7 +74,7 @@ section \Symmetry and transitivity\ Lemma (derive) pathinv: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "y =\<^bsub>A\<^esub> x" - by (equality \p:_\) intro + by (eq p) intro lemma pathinv_comp [comps]: assumes "x: A" "A: U i" @@ -70,9 +87,9 @@ Lemma (derive) pathcomp: "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" shows "x =\<^bsub>A\<^esub> z" - apply (equality \p:_\) + apply (eq p) focus prems vars x p - apply (equality \p:_\) + apply (eq p) apply intro done done @@ -82,19 +99,6 @@ lemma pathcomp_comp [comps]: shows "pathcomp A a a a (refl a) (refl a) \ refl a" unfolding pathcomp_def by reduce -text \Set up \sym\ attribute for propositional equalities:\ - -ML \ -structure Id_Sym_Attr = Sym_Attr ( - struct - val name = "sym" - val symmetry_rule = @{thm pathinv[rotated 3]} - end -) -\ - -setup \Id_Sym_Attr.setup\ - section \Notation\ @@ -113,42 +117,76 @@ translations "p \ q" \ "CONST pathcomp A x y z p q" -section \Basic propositional equalities\ +section \Calculational reasoning\ + +consts rhs :: \'a\ ("''''") + +ML \ +local fun last_rhs ctxt = + let + val this_name = Name_Space.full_name (Proof_Context.naming_of ctxt) + (Binding.name Auto_Bind.thisN) + val this = #thms (the (Proof_Context.lookup_fact ctxt this_name)) + handle Option => [] + val rhs = case map Thm.prop_of this of + [\<^const>\has_type\ $ _ $ (\<^const>\Id\ $ _ $ _ $ y)] => y + | _ => Term.dummy + in + map_aterms (fn t => case t of Const (\<^const_name>\rhs\, _) => rhs | _ => t) + end +in val _ = Context.>> + (Syntax_Phases.term_check 5 "" (fn ctxt => map (last_rhs ctxt))) +end +\ + +text \Set up \sym\ attribute for propositional equalities:\ -(*TODO: Better integration of equality type directly into reasoning...*) +ML \ +structure Id_Sym_Attr = Sym_Attr ( + struct + val name = "sym" + val symmetry_rule = @{thm pathinv[rotated 3]} + end +) +\ + +setup \Id_Sym_Attr.setup\ + + +section \Basic propositional equalities\ Lemma (derive) pathcomp_left_refl: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "(refl x) \ p = p" - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done Lemma (derive) pathcomp_right_refl: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "p \ (refl y) = p" - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done Lemma (derive) pathcomp_left_inv: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "p\ \ p = refl y" - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done Lemma (derive) pathcomp_right_inv: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "p \ p\ = refl x" - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done Lemma (derive) pathinv_pathinv: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "p\\ = p" - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done @@ -157,11 +195,11 @@ Lemma (derive) pathcomp_assoc: "A: U i" "x: A" "y: A" "z: A" "w: A" "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" "r: z =\<^bsub>A\<^esub> w" shows "p \ (q \ r) = p \ q \ r" - apply (equality \p:_\) + apply (eq p) focus prems vars x p - apply (equality \p:_\) + apply (eq p) focus prems vars x p - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done done @@ -177,7 +215,7 @@ Lemma (derive) ap: "f: A \ B" "p: x =\<^bsub>A\<^esub> y" shows "f x = f y" - apply (equality \p:_\) + apply (eq p) apply intro done @@ -199,9 +237,9 @@ Lemma (derive) ap_pathcomp: "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" shows "f[p \ q] = f[p] \ f[q]" - apply (equality \p:_\) + apply (eq p) focus prems vars x p - apply (equality \p:_\) + apply (eq p) apply (reduce; intro) done done @@ -213,7 +251,7 @@ Lemma (derive) ap_pathinv: "f: A \ B" "p: x =\<^bsub>A\<^esub> y" shows "f[p\] = f[p]\" - by (equality \p:_\) (reduce; intro) + by (eq p) (reduce; intro) text \The next two proofs currently use some low-level rewriting.\ @@ -224,7 +262,7 @@ Lemma (derive) ap_funcomp: "f: A \ B" "g: B \ C" "p: x =\<^bsub>A\<^esub> y" shows "(g \ f)[p] = g[f[p]]" - apply (equality \p:_\) + apply (eq p) apply (simp only: funcomp_apply_comp[symmetric]) apply (reduce; intro) done @@ -232,7 +270,7 @@ Lemma (derive) ap_funcomp: Lemma (derive) ap_id: assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "(id A)[p] = p" - apply (equality \p:_\) + apply (eq p) apply (rewrite at "\ = _" id_comp[symmetric]) apply (rewrite at "_ = \" id_comp[symmetric]) apply (reduce; intros) @@ -248,7 +286,7 @@ Lemma (derive) transport: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "P x \ P y" - by (equality \p:_\) intro + by (eq p) intro definition transport_i ("trans") where [implicit]: "trans P p \ transport ? P ? ? p" @@ -282,7 +320,7 @@ Lemma (derive) transport_left_inv: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "(trans P p\) \ (trans P p) = id (P x)" - by (equality \p:_\) (reduce; intro) + by (eq p) (reduce; intro) Lemma (derive) transport_right_inv: assumes @@ -291,7 +329,7 @@ Lemma (derive) transport_right_inv: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "(trans P p) \ (trans P p\) = id (P y)" - by (equality \p:_\) (reduce; intros) + by (eq p) (reduce; intros) Lemma (derive) transport_pathcomp: assumes @@ -301,9 +339,9 @@ Lemma (derive) transport_pathcomp: "u: P x" "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" shows "trans P q (trans P p u) = trans P (p \ q) u" - apply (equality \p:_\) + apply (eq p) focus prems vars x p - apply (equality \p:_\) + apply (eq p) apply (reduce; intros) done done @@ -317,7 +355,7 @@ Lemma (derive) transport_compose_typefam: "p: x =\<^bsub>A\<^esub> y" "u: P (f x)" shows "trans (\x. P (f x)) p u = trans P f[p] u" - by (equality \p:_\) (reduce; intros) + by (eq p) (reduce; intros) Lemma (derive) transport_function_family: assumes @@ -329,7 +367,7 @@ Lemma (derive) transport_function_family: "u: P x" "p: x =\<^bsub>A\<^esub> y" shows "trans Q p ((f x) u) = (f y) (trans P p u)" - by (equality \p:_\) (reduce; intros) + by (eq p) (reduce; intros) Lemma (derive) transport_const: assumes @@ -337,7 +375,7 @@ Lemma (derive) transport_const: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "\b: B. trans (\_. B) p b = b" - by (intro, equality \p:_\) (reduce; intro) + by (intro, eq p) (reduce; intro) definition transport_const_i ("trans'_const") where [implicit]: "trans_const B p \ transport_const ? B ? ? p" @@ -359,7 +397,7 @@ Lemma (derive) pathlift: "p: x =\<^bsub>A\<^esub> y" "u: P x" shows " = " - by (equality \p:_\) (reduce; intros) + by (eq p) (reduce; intros) definition pathlift_i ("lift") where [implicit]: "lift P p u \ pathlift ? P ? ? p u" @@ -383,7 +421,7 @@ Lemma (derive) pathlift_fst: "u: P x" "p: x =\<^bsub>A\<^esub> y" shows "fst[lift P p u] = p" - apply (equality \p:_\) + apply (eq p) text \Some rewriting needed here:\ \ vars x y (*Here an automatic reordering tactic would be helpful*) @@ -405,7 +443,7 @@ Lemma (derive) apd: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "trans P p (f x) = f y" - by (equality \p:_\) (reduce; intros; typechk) + by (eq p) (reduce; intros; typechk) definition apd_i ("apd") where [implicit]: "apd f p \ Identity.apd ? ? f ? ? p" @@ -428,7 +466,7 @@ Lemma (derive) apd_ap: "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" shows "apd f p = trans_const B p (f x) \ f[p]" - by (equality \p:_\) (reduce; intro) + by (eq p) (reduce; intro) end -- cgit v1.2.3