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/Equivalence.thy | 416 --------------------------------- spartan/theories/Identity.thy | 464 ------------------------------------- spartan/theories/Spartan.thy | 481 --------------------------------------- 3 files changed, 1361 deletions(-) delete mode 100644 spartan/theories/Equivalence.thy delete mode 100644 spartan/theories/Identity.thy delete mode 100644 spartan/theories/Spartan.thy (limited to 'spartan/theories') diff --git a/spartan/theories/Equivalence.thy b/spartan/theories/Equivalence.thy deleted file mode 100644 index 2975738..0000000 --- a/spartan/theories/Equivalence.thy +++ /dev/null @@ -1,416 +0,0 @@ -theory Equivalence -imports Identity - -begin - -section \Homotopy\ - -definition "homotopy A B f g \ \x: A. f `x =\<^bsub>B x\<^esub> g `x" - -definition homotopy_i (infix "~" 100) - where [implicit]: "f ~ g \ homotopy ? ? f g" - -translations "f ~ g" \ "CONST homotopy A B f g" - -Lemma homotopy_type [typechk]: - assumes - "A: U i" - "\x. x: A \ B x: U i" - "f: \x: A. B x" "g: \x: A. B x" - shows "f ~ g: U i" - unfolding homotopy_def by typechk - -Lemma (derive) homotopy_refl [refl]: - assumes - "A: U i" - "f: \x: A. B x" - shows "f ~ f" - unfolding homotopy_def by intros - -Lemma (derive) hsym: - assumes - "f: \x: A. B x" - "g: \x: A. B x" - "A: U i" - "\x. x: A \ B x: U i" - shows "H: f ~ g \ g ~ f" - unfolding homotopy_def - apply intros - apply (rule pathinv) - \ by (elim H) - \ by typechk - done - -Lemma (derive) htrans: - assumes - "f: \x: A. B x" - "g: \x: A. B x" - "h: \x: A. B x" - "A: U i" - "\x. x: A \ B x: U i" - shows "\H1: f ~ g; H2: g ~ h\ \ f ~ h" - unfolding homotopy_def - apply intro - \ vars x - apply (rule pathcomp[where ?y="g x"]) - \<^item> by (elim H1) - \<^item> by (elim H2) - done - \ by typechk - done - -text \For calculations:\ - -lemmas - homotopy_sym [sym] = hsym[rotated 4] and - homotopy_trans [trans] = htrans[rotated 5] - -Lemma (derive) commute_homotopy: - assumes - "A: U i" "B: U i" - "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" - "f: A \ B" "g: A \ B" - "H: homotopy A (\_. B) f g" - shows "(H x) \ g[p] = f[p] \ (H y)" - \ \Need this assumption unfolded for typechecking:\ - supply assms(8)[unfolded homotopy_def] - apply (eq p) - focus vars x - apply reduce - \ \Here it would really be nice to have automation for transport and - propositional equality.\ - apply (rule use_transport[where ?y="H x \ refl (g x)"]) - \ by (rule pathcomp_right_refl) - \ by (rule pathinv) (rule pathcomp_left_refl) - \ by typechk - done - done - -Corollary (derive) commute_homotopy': - assumes - "A: U i" - "x: A" - "f: A \ A" - "H: homotopy A (\_. A) f (id A)" - shows "H (f x) = f [H x]" -oops - -Lemma homotopy_id_left [typechk]: - assumes "A: U i" "B: U i" "f: A \ B" - shows "homotopy_refl A f: (id B) \ f ~ f" - unfolding homotopy_refl_def homotopy_def by reduce - -Lemma homotopy_id_right [typechk]: - assumes "A: U i" "B: U i" "f: A \ B" - shows "homotopy_refl A f: f \ (id A) ~ f" - unfolding homotopy_refl_def homotopy_def by reduce - -Lemma homotopy_funcomp_left: - assumes - "H: homotopy B C g g'" - "f: A \ B" - "g: \x: B. C x" - "g': \x: B. C x" - "A: U i" "B: U i" - "\x. x: B \ C x: U i" - shows "homotopy A {} (g \\<^bsub>A\<^esub> f) (g' \\<^bsub>A\<^esub> f)" - unfolding homotopy_def - apply (intro; reduce) - apply (insert \H: _\[unfolded homotopy_def]) - apply (elim H) - done - -Lemma homotopy_funcomp_right: - assumes - "H: homotopy A (\_. B) f f'" - "f: A \ B" - "f': A \ B" - "g: B \ C" - "A: U i" "B: U i" "C: U i" - shows "homotopy A {} (g \\<^bsub>A\<^esub> f) (g \\<^bsub>A\<^esub> f')" - unfolding homotopy_def - apply (intro; reduce) - apply (insert \H: _\[unfolded homotopy_def]) - apply (dest PiE, assumption) - apply (rule ap, assumption) - done - - -section \Quasi-inverse and bi-invertibility\ - -subsection \Quasi-inverses\ - -definition "qinv A B f \ \g: B \ A. - homotopy A (\_. A) (g \\<^bsub>A\<^esub> f) (id A) \ homotopy B (\_. B) (f \\<^bsub>B\<^esub> g) (id B)" - -lemma qinv_type [typechk]: - assumes "A: U i" "B: U i" "f: A \ B" - shows "qinv A B f: U i" - unfolding qinv_def by typechk - -definition qinv_i ("qinv") - where [implicit]: "qinv f \ Equivalence.qinv ? ? f" - -translations "qinv f" \ "CONST Equivalence.qinv A B f" - -Lemma (derive) id_qinv: - assumes "A: U i" - shows "qinv (id A)" - unfolding qinv_def - apply intro defer - apply intro defer - apply (rule homotopy_id_right) - apply (rule homotopy_id_left) - done - -Lemma (derive) quasiinv_qinv: - assumes "A: U i" "B: U i" "f: A \ B" - shows "prf: qinv f \ qinv (fst prf)" - unfolding qinv_def - apply intro - \ by (rule \f:_\) - \ apply (elim "prf") - focus vars g HH - apply intro - \<^item> by reduce (snd HH) - \<^item> by reduce (fst HH) - done - done - done - -Lemma (derive) funcomp_qinv: - assumes - "A: U i" "B: U i" "C: U i" - "f: A \ B" "g: B \ C" - shows "qinv f \ qinv g \ qinv (g \ f)" - apply (intros, unfold qinv_def, elims) - focus - prems prms - vars _ _ finv _ ginv _ HfA HfB HgB HgC - - apply intro - apply (rule funcompI[where ?f=ginv and ?g=finv]) - - text \Now a whole bunch of rewrites and we're done.\ -oops - -subsection \Bi-invertible maps\ - -definition "biinv A B f \ - (\g: B \ A. homotopy A (\_. A) (g \\<^bsub>A\<^esub> f) (id A)) - \ (\g: B \ A. homotopy B (\_. B) (f \\<^bsub>B\<^esub> g) (id B))" - -lemma biinv_type [typechk]: - assumes "A: U i" "B: U i" "f: A \ B" - shows "biinv A B f: U i" - unfolding biinv_def by typechk - -definition biinv_i ("biinv") - where [implicit]: "biinv f \ Equivalence.biinv ? ? f" - -translations "biinv f" \ "CONST Equivalence.biinv A B f" - -Lemma (derive) qinv_imp_biinv: - assumes - "A: U i" "B: U i" - "f: A \ B" - shows "?prf: qinv f \ biinv f" - apply intros - unfolding qinv_def biinv_def - by (rule Sig_dist_exp) - -text \ - Show that bi-invertible maps are quasi-inverses, as a demonstration of how to - work in this system. -\ - -Lemma (derive) biinv_imp_qinv: - assumes "A: U i" "B: U i" "f: A \ B" - shows "biinv f \ qinv f" - - text \Split the hypothesis \<^term>\biinv f\ into its components:\ - apply intro - unfolding biinv_def - apply elims - - text \Name the components:\ - focus prems vars _ _ _ g H1 h H2 - thm \g:_\ \h:_\ \H1:_\ \H2:_\ - - text \ - \<^term>\g\ is a quasi-inverse to \<^term>\f\, so the proof will be a triple - \<^term>\>\. - \ - unfolding qinv_def - apply intro - \ by (rule \g: _\) - \ apply intro - text \The first part \<^prop>\?H1: g \ f ~ id A\ is given by \<^term>\H1\.\ - apply (rule \H1: _\) - - text \ - It remains to prove \<^prop>\?H2: f \ g ~ id B\. First show that \g ~ h\, - then the result follows from @{thm \H2: f \ h ~ id B\}. Here a proof - block is used to calculate "forward". - \ - proof - - have "g ~ g \ (id B)" by reduce refl - also have "g \ (id B) ~ g \ f \ h" - by (rule homotopy_funcomp_right) (rule \H2:_\[symmetric]) - also have "g \ f \ h ~ (id A) \ h" - by (subst funcomp_assoc[symmetric]) - (rule homotopy_funcomp_left, rule \H1:_\) - also have "(id A) \ h ~ h" by reduce refl - finally have "g ~ h" by this - - then have "f \ g ~ f \ h" by (rule homotopy_funcomp_right) - - with \H2:_\ - show "f \ g ~ id B" - by (rule homotopy_trans) (assumption+, typechk) - qed - done - done - -Lemma (derive) id_biinv: - "A: U i \ biinv (id A)" - by (rule qinv_imp_biinv) (rule id_qinv) - -Lemma (derive) funcomp_biinv: - assumes - "A: U i" "B: U i" "C: U i" - "f: A \ B" "g: B \ C" - shows "biinv f \ biinv g \ biinv (g \ f)" - apply intros - focus vars biinv\<^sub>f biinv\<^sub>g - - text \Follows from \funcomp_qinv\.\ -oops - - -section \Equivalence\ - -text \ - Following the HoTT book, we first define equivalence in terms of - bi-invertibility. -\ - -definition equivalence (infix "\" 110) - where "A \ B \ \f: A \ B. Equivalence.biinv A B f" - -lemma equivalence_type [typechk]: - assumes "A: U i" "B: U i" - shows "A \ B: U i" - unfolding equivalence_def by typechk - -Lemma (derive) equivalence_refl: - assumes "A: U i" - shows "A \ A" - unfolding equivalence_def - apply intro defer - apply (rule qinv_imp_biinv) defer - apply (rule id_qinv) - done - -text \ - The following could perhaps be easier with transport (but then I think we need - univalence?)... -\ - -Lemma (derive) equivalence_symmetric: - assumes "A: U i" "B: U i" - shows "A \ B \ B \ A" - apply intros - unfolding equivalence_def - apply elim - \ vars _ f "prf" - apply (dest (4) biinv_imp_qinv) - apply intro - \<^item> unfolding qinv_def by (rule fst[of "(biinv_imp_qinv A B f) prf"]) - \<^item> by (rule qinv_imp_biinv) (rule quasiinv_qinv) - done - done - -Lemma (derive) equivalence_transitive: - assumes "A: U i" "B: U i" "C: U i" - shows "A \ B \ B \ C \ A \ C" - apply intros - unfolding equivalence_def - - text \Use \funcomp_biinv\.\ -oops - -text \ - Equal types are equivalent. We give two proofs: the first by induction, and - the second by following the HoTT book and showing that transport is an - equivalence. -\ - -Lemma - assumes - "A: U i" "B: U i" "p: A =\<^bsub>U i\<^esub> B" - shows "A \ B" - by (eq p) (rule equivalence_refl) - -text \ - The following proof is wordy because (1) the typechecker doesn't rewrite, and - (2) we don't yet have universe automation. -\ - -Lemma (derive) id_imp_equiv: - assumes - "A: U i" "B: U i" "p: A =\<^bsub>U i\<^esub> B" - shows ": A \ B" - unfolding equivalence_def - apply intros defer - - \ \Switch off auto-typechecking, which messes with universe levels\ - supply [[auto_typechk=false]] - - \ 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]) - apply (rule transport, rule U_in_U) - apply (rule lift_universe_codomain, rule U_in_U) - apply (typechk, rule U_in_U) - done - \<^item> prems vars A - apply (subst transport_comp) - \<^enum> by (rule U_in_U) - \<^enum> by reduce (rule lift_universe) - \<^enum> by reduce (rule id_biinv) - done - done - - \ \ \Similar proof as in the first subgoal above\ - apply (rewrite at A in "A \ B" id_comp[symmetric]) - apply (rewrite at B in "_ \ B" id_comp[symmetric]) - apply (rule transport, rule U_in_U) - apply (rule lift_universe_codomain, rule U_in_U) - apply (typechk, rule U_in_U) - done - done - -(*Uncomment this to see all implicits from here on. -no_translations - "f x" \ "f `x" - "x = y" \ "x =\<^bsub>A\<^esub> y" - "g \ f" \ "g \\<^bsub>A\<^esub> f" - "p\" \ "CONST pathinv A x y p" - "p \ q" \ "CONST pathcomp A x y z p q" - "fst" \ "CONST Spartan.fst A B" - "snd" \ "CONST Spartan.snd A B" - "f[p]" \ "CONST ap A B x y f p" - "trans P p" \ "CONST transport A P x y p" - "trans_const B p" \ "CONST transport_const A B x y p" - "lift P p u" \ "CONST pathlift A P x y p u" - "apd f p" \ "CONST Identity.apd A P f x y p" - "f ~ g" \ "CONST homotopy A B f g" - "qinv f" \ "CONST Equivalence.qinv A B f" - "biinv f" \ "CONST Equivalence.biinv A B f" -*) - - -end 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 diff --git a/spartan/theories/Spartan.thy b/spartan/theories/Spartan.thy deleted file mode 100644 index d235041..0000000 --- a/spartan/theories/Spartan.thy +++ /dev/null @@ -1,481 +0,0 @@ -text \Spartan type theory\ - -theory Spartan -imports - Pure - "HOL-Eisbach.Eisbach" - "HOL-Eisbach.Eisbach_Tools" -keywords - "Theorem" "Lemma" "Corollary" "Proposition" :: thy_goal_stmt and - "focus" "\" "\<^item>" "\<^enum>" "~" :: prf_script_goal % "proof" and - "derive" "prems" "vars":: quasi_command and - "print_coercions" :: thy_decl - -begin - - -section \Preamble\ - -declare [[eta_contract=false]] - - -section \Metatype setup\ - -typedecl o - - -section \Judgments\ - -judgment has_type :: \o \ o \ prop\ ("(2_:/ _)" 999) - - -section \Universes\ - -typedecl lvl \ \Universe levels\ - -axiomatization - O :: \lvl\ and - S :: \lvl \ lvl\ and - lt :: \lvl \ lvl \ prop\ (infix "<" 900) - where - O_min: "O < S i" and - lt_S: "i < S i" and - lt_trans: "i < j \ j < k \ i < k" - -axiomatization U :: \lvl \ o\ where - U_hierarchy: "i < j \ U i: U j" and - U_cumulative: "A: U i \ i < j \ A: U j" - -lemma U_in_U: - "U i: U (S i)" - by (rule U_hierarchy, rule lt_S) - -lemma lift_universe: - "A: U i \ A: U (S i)" - by (erule U_cumulative, rule lt_S) - - -section \\-type\ - -axiomatization - Pi :: \o \ (o \ o) \ o\ and - lam :: \o \ (o \ o) \ o\ and - app :: \o \ o \ o\ ("(1_ `_)" [120, 121] 120) - -syntax - "_Pi" :: \idt \ o \ o \ o\ ("(2\_: _./ _)" 30) - "_lam" :: \pttrns \ o \ o \ o\ ("(2\_: _./ _)" 30) - "_lam2" :: \pttrns \ o \ o \ o\ -translations - "\x: A. B" \ "CONST Pi A (\x. B)" - "\x xs: A. b" \ "CONST lam A (\x. _lam2 xs A b)" - "_lam2 x A b" \ "\x: A. b" - "\x: A. b" \ "CONST lam A (\x. b)" - -abbreviation Fn (infixr "\" 40) where "A \ B \ \_: A. B" - -axiomatization where - PiF: "\\x. x: A \ B x: U i; A: U i\ \ \x: A. B x: U i" and - - PiI: "\\x. x: A \ b x: B x; A: U i\ \ \x: A. b x: \x: A. B x" and - - PiE: "\f: \x: A. B x; a: A\ \ f `a: B a" and - - beta: "\a: A; \x. x: A \ b x: B x\ \ (\x: A. b x) `a \ b a" and - - eta: "f: \x: A. B x \ \x: A. f `x \ f" and - - Pi_cong: "\ - A: U i; - \x. x: A \ B x: U i; - \x. x: A \ B' x: U i; - \x. x: A \ B x \ B' x - \ \ \x: A. B x \ \x: A. B' x" and - - lam_cong: "\\x. x: A \ b x \ c x; A: U i\ \ \x: A. b x \ \x: A. c x" - - -section \\-type\ - -axiomatization - Sig :: \o \ (o \ o) \ o\ and - pair :: \o \ o \ o\ ("(2<_,/ _>)") and - SigInd :: \o \ (o \ o) \ (o \ o) \ (o \ o \ o) \ o \ o\ - -syntax "_Sum" :: \idt \ o \ o \ o\ ("(2\_: _./ _)" 20) - -translations "\x: A. B" \ "CONST Sig A (\x. B)" - -abbreviation Prod (infixl "\" 50) - where "A \ B \ \_: A. B" - -axiomatization where - SigF: "\\x. x: A \ B x: U i; A: U i\ \ \x: A. B x: U i" and - - SigI: "\\x. x: A \ B x: U i; a: A; b: B a\ \ : \x: A. B x" and - - SigE: "\ - p: \x: A. B x; - A: U i; - \x. x : A \ B x: U i; - \p. p: \x: A. B x \ C p: U i; - \x y. \x: A; y: B x\ \ f x y: C - \ \ SigInd A (\x. B x) (\p. C p) f p: C p" and - - Sig_comp: "\ - a: A; - b: B a; - \x. x: A \ B x: U i; - \p. p: \x: A. B x \ C p: U i; - \x y. \x: A; y: B x\ \ f x y: C - \ \ SigInd A (\x. B x) (\p. C p) f \ f a b" and - - Sig_cong: "\ - \x. x: A \ B x \ B' x; - A: U i; - \x. x : A \ B x: U i; - \x. x : A \ B' x: U i - \ \ \x: A. B x \ \x: A. B' x" - - -section \Proof commands\ - -named_theorems typechk - -ML_file \../lib/lib.ML\ -ML_file \../lib/goals.ML\ -ML_file \../lib/focus.ML\ - - -section \Congruence automation\ - -(*Potential to be retired*) -ML_file \../lib/congruence.ML\ - - -section \Methods\ - -ML_file \../lib/elimination.ML\ \ \declares the [elims] attribute\ - -named_theorems intros and comps -lemmas - [intros] = PiF PiI SigF SigI and - [elims "?f"] = PiE and - [elims "?p"] = SigE and - [comps] = beta Sig_comp and - [cong] = Pi_cong lam_cong Sig_cong - -ML_file \../lib/tactics.ML\ - -method_setup assumptions = - \Scan.succeed (fn ctxt => SIMPLE_METHOD ( - CHANGED (TRYALL (assumptions_tac ctxt))))\ - -method_setup known = - \Scan.succeed (fn ctxt => SIMPLE_METHOD ( - CHANGED (TRYALL (known_tac ctxt))))\ - -method_setup intro = - \Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (intro_tac ctxt)))\ - -method_setup intros = - \Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (intros_tac ctxt)))\ - -method_setup old_elim = - \Scan.option Args.term >> (fn tm => fn ctxt => - SIMPLE_METHOD' (SIDE_CONDS (old_elims_tac tm ctxt) ctxt))\ - -method_setup elim = - \Scan.repeat Args.term >> (fn tms => fn ctxt => - CONTEXT_METHOD (K (elim_context_tac tms ctxt 1)))\ - -method elims = elim+ - -(*This could possibly use additional simplification hints via a simp: modifier*) -method_setup typechk = - \Scan.succeed (fn ctxt => SIMPLE_METHOD' ( - SIDE_CONDS (typechk_tac ctxt) ctxt)) - (* CHANGED (ALLGOALS (TRY o typechk_tac ctxt)))) *)\ - -method_setup rule = - \Attrib.thms >> (fn ths => fn ctxt => - SIMPLE_METHOD (HEADGOAL (SIDE_CONDS (rule_tac ths ctxt) ctxt)))\ - -method_setup dest = - \Scan.lift (Scan.option (Args.parens Parse.int)) -- Attrib.thms - >> (fn (opt_n, ths) => fn ctxt => - SIMPLE_METHOD (HEADGOAL (SIDE_CONDS (dest_tac opt_n ths ctxt) ctxt)))\ - -subsection \Reflexivity\ - -named_theorems refl -method refl = (rule refl) - -subsection \Trivial proofs modulo typechecking\ - -method_setup this = - \Scan.succeed (fn ctxt => METHOD ( - EVERY o map (HEADGOAL o - (fn ths => SIDE_CONDS (resolve_tac ctxt ths) ctxt) o - single) - ))\ - -subsection \Rewriting\ - -\ \\subst\ method\ -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 "../lib/eqsubst.ML" - -\ \\rewrite\ method\ -consts rewrite_HOLE :: "'a::{}" ("\") - -lemma eta_expand: - fixes f :: "'a::{} \ 'b::{}" - shows "f \ \x. f x" . - -lemma rewr_imp: - assumes "PROP A \ PROP B" - shows "(PROP A \ PROP C) \ (PROP B \ PROP C)" - apply (rule Pure.equal_intr_rule) - apply (drule equal_elim_rule2[OF assms]; assumption) - apply (drule equal_elim_rule1[OF assms]; assumption) - done - -lemma imp_cong_eq: - "(PROP A \ (PROP B \ PROP C) \ (PROP B' \ PROP C')) \ - ((PROP B \ PROP A \ PROP C) \ (PROP B' \ PROP A \ PROP C'))" - apply (Pure.intro Pure.equal_intr_rule) - apply (drule (1) cut_rl; drule Pure.equal_elim_rule1 Pure.equal_elim_rule2; - assumption)+ - apply (drule Pure.equal_elim_rule1 Pure.equal_elim_rule2; assumption)+ - done - -ML_file \~~/src/HOL/Library/cconv.ML\ -ML_file \../lib/rewrite.ML\ - -\ \\reduce\ computes terms via judgmental equalities\ -setup \map_theory_simpset (fn ctxt => ctxt addSolver (mk_solver "" typechk_tac))\ - -method reduce uses add = (simp add: comps add | subst comps)+ - - -section \Implicit notations\ - -text \ - \?\ is used to mark implicit arguments in definitions, while \{}\ is expanded - immediately for elaboration in statements. -\ - -consts - iarg :: \'a\ ("?") - hole :: \'b\ ("{}") - -ML_file \../lib/implicits.ML\ - -attribute_setup implicit = \Scan.succeed Implicits.implicit_defs_attr\ - -ML \ -val _ = Context.>> - (Syntax_Phases.term_check 1 "" (fn ctxt => map (Implicits.make_holes ctxt))) -\ - -text \Automatically insert inhabitation judgments where needed:\ - -consts inhabited :: \o \ prop\ ("(_)") -translations "CONST inhabited A" \ "CONST has_type {} A" - - -section \Lambda coercion\ - -\ \Coerce object lambdas to meta-lambdas\ -abbreviation (input) lambda :: \o \ o \ o\ - where "lambda f \ \x. f `x" - -ML_file \~~/src/Tools/subtyping.ML\ -declare [[coercion_enabled, coercion lambda]] - -translations "f x" \ "f `x" - - -section \Functions\ - -lemma eta_exp: - assumes "f: \x: A. B x" - shows "f \ \x: A. f x" - by (rule eta[symmetric]) - -lemma lift_universe_codomain: - assumes "A: U i" "f: A \ U j" - shows "f: A \ U (S j)" - apply (sub eta_exp) - apply known - apply (Pure.rule intros; rule lift_universe) - done - -subsection \Function composition\ - -definition "funcomp A g f \ \x: A. g `(f `x)" - -syntax - "_funcomp" :: \o \ o \ o \ o\ ("(2_ \\<^bsub>_\<^esub>/ _)" [111, 0, 110] 110) -translations - "g \\<^bsub>A\<^esub> f" \ "CONST funcomp A g f" - -lemma funcompI [typechk]: - assumes - "A: U i" - "B: U i" - "\x. x: B \ C x: U i" - "f: A \ B" - "g: \x: B. C x" - shows - "g \\<^bsub>A\<^esub> f: \x: A. C (f x)" - unfolding funcomp_def by typechk - -lemma funcomp_assoc [comps]: - assumes - "f: A \ B" - "g: B \ C" - "h: \x: C. D x" - "A: U i" - shows - "(h \\<^bsub>B\<^esub> g) \\<^bsub>A\<^esub> f \ h \\<^bsub>A\<^esub> g \\<^bsub>A\<^esub> f" - unfolding funcomp_def by reduce - -lemma funcomp_lambda_comp [comps]: - assumes - "A: U i" - "\x. x: A \ b x: B" - "\x. x: B \ c x: C x" - shows - "(\x: B. c x) \\<^bsub>A\<^esub> (\x: A. b x) \ \x: A. c (b x)" - unfolding funcomp_def by reduce - -lemma funcomp_apply_comp [comps]: - assumes - "f: A \ B" "g: \x: B. C x" - "x: A" - "A: U i" "B: U i" - "\x y. x: B \ C x: U i" - shows "(g \\<^bsub>A\<^esub> f) x \ g (f x)" - unfolding funcomp_def by reduce - -text \Notation:\ - -definition funcomp_i (infixr "\" 120) - where [implicit]: "funcomp_i g f \ g \\<^bsub>?\<^esub> f" - -translations "g \ f" \ "g \\<^bsub>A\<^esub> f" - -subsection \Identity function\ - -abbreviation id where "id A \ \x: A. x" - -lemma - id_type[typechk]: "A: U i \ id A: A \ A" and - id_comp [comps]: "x: A \ (id A) x \ x" \ \for the occasional manual rewrite\ - by reduce - -lemma id_left [comps]: - assumes "f: A \ B" "A: U i" "B: U i" - shows "(id B) \\<^bsub>A\<^esub> f \ f" - by (subst eta_exp[of f]) (reduce, rule eta) - -lemma id_right [comps]: - assumes "f: A \ B" "A: U i" "B: U i" - shows "f \\<^bsub>A\<^esub> (id A) \ f" - by (subst eta_exp[of f]) (reduce, rule eta) - -lemma id_U [typechk]: - "id (U i): U i \ U i" - by typechk (fact U_in_U) - - -section \Pairs\ - -definition "fst A B \ \p: \x: A. B x. SigInd A B (\_. A) (\x y. x) p" -definition "snd A B \ \p: \x: A. B x. SigInd A B (\p. B (fst A B p)) (\x y. y) p" - -lemma fst_type [typechk]: - assumes "A: U i" "\x. x: A \ B x: U i" - shows "fst A B: (\x: A. B x) \ A" - unfolding fst_def by typechk - -lemma fst_comp [comps]: - assumes - "a: A" - "b: B a" - "A: U i" - "\x. x: A \ B x: U i" - shows "fst A B \ a" - unfolding fst_def by reduce - -lemma snd_type [typechk]: - assumes "A: U i" "\x. x: A \ B x: U i" - shows "snd A B: \p: \x: A. B x. B (fst A B p)" - unfolding snd_def by typechk reduce - -lemma snd_comp [comps]: - assumes "a: A" "b: B a" "A: U i" "\x. x: A \ B x: U i" - shows "snd A B \ b" - unfolding snd_def by reduce - -subsection \Notation\ - -definition fst_i ("fst") - where [implicit]: "fst \ Spartan.fst ? ?" - -definition snd_i ("snd") - where [implicit]: "snd \ Spartan.snd ? ?" - -translations - "fst" \ "CONST Spartan.fst A B" - "snd" \ "CONST Spartan.snd A B" - -subsection \Projections\ - -Lemma fst [typechk]: - assumes - "p: \x: A. B x" - "A: U i" "\x. x: A \ B x: U i" - shows "fst p: A" - by typechk - -Lemma snd [typechk]: - assumes - "p: \x: A. B x" - "A: U i" "\x. x: A \ B x: U i" - shows "snd p: B (fst p)" - by typechk - -method fst for p::o = rule fst[of p] -method snd for p::o = rule snd[of p] - -subsection \Properties of \\ - -Lemma (derive) Sig_dist_exp: - assumes - "p: \x: A. B x \ C x" - "A: U i" - "\x. x: A \ B x: U i" - "\x. x: A \ C x: U i" - shows "(\x: A. B x) \ (\x: A. C x)" - apply (elim p) - focus vars x y - apply intro - \ apply intro - apply assumption - apply (fst y) - done - \ apply intro - apply assumption - apply (snd y) - done - done - done - - -end -- cgit v1.2.3