From 80edbd08e13200d2c080ac281d19948bbbcd92e0 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 25 May 2020 17:09:54 +0200 Subject: Reorganize theory structure. In particular, the identity type moves out from under Spartan to HoTT. Spartan now only has Pi and Sigma. --- spartan/theories/Identity.thy | 464 ------------------------------------------ 1 file changed, 464 deletions(-) delete mode 100644 spartan/theories/Identity.thy (limited to 'spartan/theories/Identity.thy') diff --git a/spartan/theories/Identity.thy b/spartan/theories/Identity.thy deleted file mode 100644 index 3a982f6..0000000 --- a/spartan/theories/Identity.thy +++ /dev/null @@ -1,464 +0,0 @@ -chapter \The identity type\ - -text \ - The identity type, the higher groupoid structure of types, - and type families as fibrations. -\ - -theory Identity -imports Spartan - -begin - -axiomatization - Id :: \o \ o \ o \ o\ and - refl :: \o \ o\ and - IdInd :: \o \ (o \ o \ o \ o) \ (o \ o) \ o \ o \ o \ o\ - -syntax "_Id" :: \o \ o \ o \ o\ ("(2_ =\<^bsub>_\<^esub>/ _)" [111, 0, 111] 110) - -translations "a =\<^bsub>A\<^esub> b" \ "CONST Id A a b" - -axiomatization where - IdF: "\A: U i; a: A; b: A\ \ a =\<^bsub>A\<^esub> b: U i" and - - IdI: "a: A \ refl a: a =\<^bsub>A\<^esub> a" and - - IdE: "\ - p: a =\<^bsub>A\<^esub> b; - a: A; - b: A; - \x y p. \p: x =\<^bsub>A\<^esub> y; x: A; y: A\ \ C x y p: U i; - \x. x: A \ f x: C x x (refl x) - \ \ IdInd A (\x y p. C x y p) f a b p: C a b p" and - - Id_comp: "\ - a: A; - \x y p. \x: A; y: A; p: x =\<^bsub>A\<^esub> y\ \ C x y p: U i; - \x. x: A \ f x: C x x (refl x) - \ \ IdInd A (\x y p. C x y p) f a a (refl a) \ f a" - -lemmas - [intros] = IdF IdI and - [elims "?p" "?a" "?b"] = IdE and - [comps] = Id_comp and - [refl] = IdI - - -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\ - -Lemma (derive) pathinv: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" - shows "y =\<^bsub>A\<^esub> x" - by (eq p) intro - -lemma pathinv_comp [comps]: - assumes "x: A" "A: U i" - shows "pathinv A x x (refl x) \ refl x" - unfolding pathinv_def by reduce - -Lemma (derive) pathcomp: - assumes - "A: U i" "x: A" "y: A" "z: A" - "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" - shows - "x =\<^bsub>A\<^esub> z" - apply (eq p) - focus prems vars x p - apply (eq p) - apply intro - done - done - -lemma pathcomp_comp [comps]: - assumes "a: A" "A: U i" - shows "pathcomp A a a a (refl a) (refl a) \ refl a" - unfolding pathcomp_def by reduce - - -section \Notation\ - -definition Id_i (infix "=" 110) - where [implicit]: "Id_i x y \ x =\<^bsub>?\<^esub> y" - -definition pathinv_i ("_\" [1000]) - where [implicit]: "pathinv_i p \ pathinv ? ? ? p" - -definition pathcomp_i (infixl "\" 121) - where [implicit]: "pathcomp_i p q \ pathcomp ? ? ? ? p q" - -translations - "x = y" \ "x =\<^bsub>A\<^esub> y" - "p\" \ "CONST pathinv A x y p" - "p \ q" \ "CONST pathcomp A x y z p q" - - -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 -\ - -lemmas - [sym] = pathinv[rotated 3] and - [trans] = pathcomp[rotated 4] - - -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 (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 (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 (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 (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 (eq p) - apply (reduce; intros) - done - -Lemma (derive) pathcomp_assoc: - assumes - "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 (eq p) - focus prems vars x p - apply (eq p) - focus prems vars x p - apply (eq p) - apply (reduce; intros) - done - done - done - - -section \Functoriality of functions\ - -Lemma (derive) ap: - assumes - "A: U i" "B: U i" - "x: A" "y: A" - "f: A \ B" - "p: x =\<^bsub>A\<^esub> y" - shows "f x = f y" - apply (eq p) - apply intro - done - -definition ap_i ("_[_]" [1000, 0]) - where [implicit]: "ap_i f p \ ap ? ? ? ? f p" - -translations "f[p]" \ "CONST ap A B x y f p" - -Lemma ap_refl [comps]: - assumes "f: A \ B" "x: A" "A: U i" "B: U i" - shows "f[refl x] \ refl (f x)" - unfolding ap_def by reduce - -Lemma (derive) ap_pathcomp: - assumes - "A: U i" "B: U i" - "x: A" "y: A" "z: A" - "f: A \ B" - "p: x =\<^bsub>A\<^esub> y" "q: y =\<^bsub>A\<^esub> z" - shows - "f[p \ q] = f[p] \ f[q]" - apply (eq p) - focus prems vars x p - apply (eq p) - apply (reduce; intro) - done - done - -Lemma (derive) ap_pathinv: - assumes - "A: U i" "B: U i" - "x: A" "y: A" - "f: A \ B" - "p: x =\<^bsub>A\<^esub> y" - shows "f[p\] = f[p]\" - by (eq p) (reduce; intro) - -text \The next two proofs currently use some low-level rewriting.\ - -Lemma (derive) ap_funcomp: - assumes - "A: U i" "B: U i" "C: U i" - "x: A" "y: A" - "f: A \ B" "g: B \ C" - "p: x =\<^bsub>A\<^esub> y" - shows "(g \ f)[p] = g[f[p]]" - apply (eq p) - apply (simp only: funcomp_apply_comp[symmetric]) - apply (reduce; intro) - done - -Lemma (derive) ap_id: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" - shows "(id A)[p] = p" - apply (eq p) - apply (rewrite at "\ = _" id_comp[symmetric]) - apply (rewrite at "_ = \" id_comp[symmetric]) - apply (reduce; intros) - done - - -section \Transport\ - -Lemma (derive) transport: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "P x \ P y" - by (eq p) intro - -definition transport_i ("trans") - where [implicit]: "trans P p \ transport ? P ? ? p" - -translations "trans P p" \ "CONST transport A P x y p" - -Lemma transport_comp [comps]: - assumes - "a: A" - "A: U i" - "\x. x: A \ P x: U i" - shows "trans P (refl a) \ id (P a)" - unfolding transport_def by reduce - -\ \TODO: Build transport automation!\ - -Lemma use_transport: - assumes - "p: y =\<^bsub>A\<^esub> x" - "u: P x" - "x: A" "y: A" - "A: U i" - "\x. x: A \ P x: U i" - shows "trans P p\ u: P y" - by typechk - -Lemma (derive) transport_left_inv: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "(trans P p\) \ (trans P p) = id (P x)" - by (eq p) (reduce; intro) - -Lemma (derive) transport_right_inv: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "(trans P p) \ (trans P p\) = id (P y)" - by (eq p) (reduce; intros) - -Lemma (derive) transport_pathcomp: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" "z: A" - "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 (eq p) - focus prems vars x p - apply (eq p) - apply (reduce; intros) - done - done - -Lemma (derive) transport_compose_typefam: - assumes - "A: U i" "B: U i" - "\x. x: B \ P x: U i" - "f: A \ B" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - "u: P (f x)" - shows "trans (\x. P (f x)) p u = trans P f[p] u" - by (eq p) (reduce; intros) - -Lemma (derive) transport_function_family: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "\x. x: A \ Q x: U i" - "f: \x: A. P x \ Q x" - "x: A" "y: A" - "u: P x" - "p: x =\<^bsub>A\<^esub> y" - shows "trans Q p ((f x) u) = (f y) (trans P p u)" - by (eq p) (reduce; intros) - -Lemma (derive) transport_const: - assumes - "A: U i" "B: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "\b: B. trans (\_. B) p b = b" - by (intro, eq p) (reduce; intro) - -definition transport_const_i ("trans'_const") - where [implicit]: "trans_const B p \ transport_const ? B ? ? p" - -translations "trans_const B p" \ "CONST transport_const A B x y p" - -Lemma transport_const_comp [comps]: - assumes - "x: A" "b: B" - "A: U i" "B: U i" - shows "trans_const B (refl x) b\ refl b" - unfolding transport_const_def by reduce - -Lemma (derive) pathlift: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - "u: P x" - shows " = " - by (eq p) (reduce; intros) - -definition pathlift_i ("lift") - where [implicit]: "lift P p u \ pathlift ? P ? ? p u" - -translations "lift P p u" \ "CONST pathlift A P x y p u" - -Lemma pathlift_comp [comps]: - assumes - "u: P x" - "x: A" - "\x. x: A \ P x: U i" - "A: U i" - shows "lift P (refl x) u \ refl " - unfolding pathlift_def by reduce - -Lemma (derive) pathlift_fst: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "x: A" "y: A" - "u: P x" - "p: x =\<^bsub>A\<^esub> y" - shows "fst[lift P p u] = p" - apply (eq p) - text \Some rewriting needed here:\ - \ vars x y - (*Here an automatic reordering tactic would be helpful*) - apply (rewrite at x in "x = y" fst_comp[symmetric]) - prefer 4 - apply (rewrite at y in "_ = y" fst_comp[symmetric]) - done - \ by reduce intro - done - - -section \Dependent paths\ - -Lemma (derive) apd: - assumes - "A: U i" - "\x. x: A \ P x: U i" - "f: \x: A. P x" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "trans P p (f x) = f y" - by (eq p) (reduce; intros; typechk) - -definition apd_i ("apd") - where [implicit]: "apd f p \ Identity.apd ? ? f ? ? p" - -translations "apd f p" \ "CONST Identity.apd A P f x y p" - -Lemma dependent_map_comp [comps]: - assumes - "f: \x: A. P x" - "x: A" - "A: U i" - "\x. x: A \ P x: U i" - shows "apd f (refl x) \ refl (f x)" - unfolding apd_def by reduce - -Lemma (derive) apd_ap: - assumes - "A: U i" "B: U i" - "f: A \ B" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - shows "apd f p = trans_const B p (f x) \ f[p]" - by (eq p) (reduce; intro) - - -end -- cgit v1.2.3