From 0f84aa06384d827bd5f5f137c218197ab7217762 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 27 Jul 2020 14:29:48 +0200 Subject: Hook elaboration into assumptions mechanism --- hott/Identity.thy | 82 ++++++++++++++++++++++++++++--------------------------- 1 file changed, 42 insertions(+), 40 deletions(-) (limited to 'hott/Identity.thy') diff --git a/hott/Identity.thy b/hott/Identity.thy index 29ce26a..728537c 100644 --- a/hott/Identity.thy +++ b/hott/Identity.thy @@ -135,14 +135,14 @@ lemmas section \Basic propositional equalities\ Lemma (derive) refl_pathcomp: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "(refl x) \ p = p" apply (eq p) apply (reduce; intro) done Lemma (derive) pathcomp_refl: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "p \ (refl y) = p" apply (eq p) apply (reduce; intro) @@ -166,24 +166,24 @@ Lemma ru_refl [comp]: unfolding pathcomp_refl_def by reduce Lemma (derive) inv_pathcomp: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "p\ \ p = refl y" by (eq p) (reduce; intro) Lemma (derive) pathcomp_inv: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "p \ p\ = refl x" by (eq p) (reduce; intro) Lemma (derive) pathinv_pathinv: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "p\\ = p" by (eq p) (reduce; intro) 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" + "p: x = y" "q: y = z" "r: z = w" shows "p \ (q \ r) = p \ q \ r" apply (eq p) focus prems vars x p @@ -203,7 +203,7 @@ Lemma (derive) ap: "A: U i" "B: U i" "x: A" "y: A" "f: A \ B" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "f x = f y" by (eq p) intro @@ -222,7 +222,7 @@ Lemma (derive) ap_pathcomp: "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" + "p: x = y" "q: y = z" shows "f[p \ q] = f[p] \ f[q]" apply (eq p) @@ -237,7 +237,7 @@ Lemma (derive) ap_pathinv: "A: U i" "B: U i" "x: A" "y: A" "f: A \ B" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "f[p\] = f[p]\" by (eq p) (reduce; intro) @@ -248,7 +248,7 @@ Lemma (derive) ap_funcomp: "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" + "p: x = y" shows "(g \ f)[p] = g[f[p]]" apply (eq p) \ by reduce @@ -256,7 +256,7 @@ Lemma (derive) ap_funcomp: done Lemma (derive) ap_id: - assumes "A: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" + assumes "A: U i" "x: A" "y: A" "p: x = y" shows "(id A)[p] = p" apply (eq p) \ by reduce @@ -271,7 +271,7 @@ Lemma (derive) transport: "A: U i" "\x. x: A \ P x: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "P x \ P y" by (eq p) intro @@ -308,7 +308,7 @@ Lemma (derive) transport_left_inv: "A: U i" "\x. x: A \ P x: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "(trans P p\) \ (trans P p) = id (P x)" by (eq p) (reduce; refl) @@ -317,7 +317,7 @@ Lemma (derive) transport_right_inv: "A: U i" "\x. x: A \ P x: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "(trans P p) \ (trans P p\) = id (P y)" by (eq p) (reduce; intro) @@ -327,7 +327,7 @@ Lemma (derive) transport_pathcomp: "\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" + "p: x = y" "q: y = z" shows "trans P q (trans P p u) = trans P (p \ q) u" apply (eq p) focus prems vars x p @@ -342,7 +342,7 @@ Lemma (derive) transport_compose_typefam: "\x. x: B \ P x: U i" "f: A \ B" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" "u: P (f x)" shows "trans (fn x. P (f x)) p u = trans P f[p] u" by (eq p) (reduce; intro) @@ -355,7 +355,7 @@ Lemma (derive) transport_function_family: "f: \x: A. P x \ Q x" "x: A" "y: A" "u: P x" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "trans Q p ((f x) u) = (f y) (trans P p u)" by (eq p) (reduce; intro) @@ -363,7 +363,7 @@ Lemma (derive) transport_const: assumes "A: U i" "B: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "\b: B. trans (fn _. B) p b = b" by (intro, eq p) (reduce; intro) @@ -384,7 +384,7 @@ Lemma (derive) pathlift: "A: U i" "\x. x: A \ P x: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" "u: P x" shows " = " by (eq p) (reduce; intro) @@ -409,7 +409,7 @@ Lemma (derive) pathlift_fst: "\x. x: A \ P x: U i" "x: A" "y: A" "u: P x" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "fst[lift P p u] = p" apply (eq p) \ by reduce @@ -425,7 +425,7 @@ Lemma (derive) apd: "\x. x: A \ P x: U i" "f: \x: A. P x" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "trans P p (f x) = f y" by (eq p) (reduce; intro) @@ -448,7 +448,7 @@ Lemma (derive) apd_ap: "A: U i" "B: U i" "f: A \ B" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" shows "apd f p = trans_const B p (f x) \ f[p]" by (eq p) (reduce; intro) @@ -457,10 +457,9 @@ section \Whiskering\ Lemma (derive) right_whisker: assumes "A: U i" "a: A" "b: A" "c: A" - shows "\p: a = b; q: a = b; r: b = c; \: p =\<^bsub>a = b\<^esub> q\ \ p \ r = q \ r" - \ \TODO: In the above we need to annotate the type of \ with the type `a = b` - in order for the `eq` method to work correctly. This should be fixed with a - pre-proof elaborator.\ + and "p: a = b" "q: a = b" "r: b = c" + and "\: p = q" + shows "p \ r = q \ r" apply (eq r) focus prems vars x s t proof - @@ -473,7 +472,9 @@ Lemma (derive) right_whisker: Lemma (derive) left_whisker: assumes "A: U i" "a: A" "b: A" "c: A" - shows "\p: b = c; q: b = c; r: a = b; \: p =\<^bsub>b = c\<^esub> q\ \ r \ p = r \ q" + and "p: b = c" "q: b = c" "r: a = b" + and "\: p = q" + shows "r \ p = r \ q" apply (eq r) focus prems prms vars x s t proof - @@ -495,15 +496,13 @@ translations "r \\<^sub>l\<^bsub>c\<^esub> \" \ "CONST left_whisker A a b c p q r \" Lemma whisker_refl [comp]: - assumes "A: U i" "a: A" "b: A" - shows "\p: a = b; q: a = b; \: p =\<^bsub>a = b\<^esub> q\ \ - \ \\<^sub>r\<^bsub>a\<^esub> (refl b) \ ru p \ \ \ (ru q)\" + assumes "A: U i" "a: A" "b: A" "p: a = b" "q: a = b" "\: p = q" + shows "\ \\<^sub>r\<^bsub>a\<^esub> (refl b) \ ru p \ \ \ (ru q)\" unfolding right_whisker_def by reduce Lemma refl_whisker [comp]: - assumes "A: U i" "a: A" "b: A" - shows "\p: a = b; q: a = b; \: p = q\ \ - (refl a) \\<^sub>l\<^bsub>b\<^esub> \ \ (lu p) \ \ \ (lu q)\" + assumes "A: U i" "a: A" "b: A" "p: a = b" "q: a = b" "\: p = q" + shows "(refl a) \\<^sub>l\<^bsub>b\<^esub> \ \ (lu p) \ \ \ (lu q)\" unfolding left_whisker_def by reduce method left_whisker = (rule left_whisker) @@ -524,20 +523,22 @@ begin Lemma (derive) horiz_pathcomp: notes assums - shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" + assumes "\: p = q" "\: r = s" + shows "p \ r = q \ s" proof (rule pathcomp) - show "\: p = q \ p \ r = q \ r" by right_whisker - show "\: r = s \ .. = q \ s" by left_whisker + show "p \ r = q \ r" by right_whisker fact + show ".. = q \ s" by left_whisker fact qed typechk text \A second horizontal composition:\ Lemma (derive) horiz_pathcomp': notes assums - shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" + assumes "\: p = q" "\: r = s" + shows "p \ r = q \ s" proof (rule pathcomp) - show "\: r = s \ p \ r = p \ s" by left_whisker - show "\: p = q \ .. = q \ s" by right_whisker + show "p \ r = p \ s" by left_whisker fact + show ".. = q \ s" by right_whisker fact qed typechk notation horiz_pathcomp (infix "\" 121) @@ -545,7 +546,8 @@ notation horiz_pathcomp' (infix "\''" 121) Lemma (derive) horiz_pathcomp_eq_horiz_pathcomp': notes assums - shows "\\: p = q; \: r = s\ \ \ \ \ = \ \' \" + assumes "\: p = q" "\: r = s" + shows "\ \ \ = \ \' \" unfolding horiz_pathcomp_def horiz_pathcomp'_def apply (eq \, eq \) focus vars p apply (eq p) -- cgit v1.2.3