From 7a7e27f4a1efd69e9ab43b95d3c7ead61a743e55 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 4 Feb 2019 11:53:47 +0100 Subject: 1. Change syntax to rely less on unicode/control symbols. 2. Begin work on object-level type inference and input syntax help. --- Prod.thy | 101 ++++++++++++++++++++++++++++++++++++++------------------------- 1 file changed, 61 insertions(+), 40 deletions(-) (limited to 'Prod.thy') diff --git a/Prod.thy b/Prod.thy index 8d840bd..68c9405 100644 --- a/Prod.thy +++ b/Prod.thy @@ -1,91 +1,112 @@ -(* -Title: Prod.thy -Author: Joshua Chen -Date: 2018 +(******** +Isabelle/HoTT: Dependent product (dependent function) +Jan 2019 -Dependent product type -*) +********) theory Prod imports HoTT_Base HoTT_Methods begin - -section \Basic definitions\ +section \Basic type definitions\ axiomatization - Prod :: "[t, tf] \ t" and - lambda :: "(t \ t) \ t" (binder "\<^bold>\" 30) and - appl :: "[t, t] \ t" ("(1_`_)" [120, 121] 120) \ \Application binds tighter than abstraction.\ + Prod :: "[t, t \ t] \ t" and + lam :: "[t, t \ t] \ t" and + app :: "[t, t] \ t" ("(1_ ` _)" [120, 121] 120) + \ \Application should bind tighter than abstraction.\ syntax - "_prod" :: "[idt, t, t] \ t" ("(3\_: _./ _)" 30) - "_prod_ascii" :: "[idt, t, t] \ t" ("(3II _: _./ _)" 30) - -text \The translations below bind the variable @{term x} in the expressions @{term B} and @{term b}.\ - + "_Prod" :: "[idt, t, t] \ t" ("(3TT '(_: _')./ _)" 30) + "_Prod'" :: "[idt, t, t] \ t" ("(3TT _: _./ _)" 30) + "_lam" :: "[idt, t, t] \ t" ("(3,\\ '(_: _')./ _)" 30) + "_lam'" :: "[idt, t, t] \ t" ("(3,\\ _: _./ _)" 30) translations - "\x:A. B" \ "CONST Prod A (\x. B)" - "II x:A. B" \ "CONST Prod A (\x. B)" + "TT(x: A). B" \ "(CONST Prod) A (\x. B)" + "TT x: A. B" \ "(CONST Prod) A (\x. B)" + ",\\(x: A). b" \ "(CONST lam) A (\x. b)" + ",\\x: A. b" \ "(CONST lam) A (\x. b)" + +text \ +The syntax translations above bind the variable @{term x} in the expressions @{term B} and @{term b}. +\ -text \Non-dependent functions are a special case.\ +text \Non-dependent functions are a special case:\ abbreviation Fun :: "[t, t] \ t" (infixr "\" 40) - where "A \ B \ \_: A. B" +where "A \ B \ TT(_: A). B" axiomatization where \ \Type rules\ - Prod_form: "\A: U i; B: A \ U i\ \ \x:A. B x: U i" and + Prod_form: "\A: U i; \x. x: A \ B x: U i\ \ TT x: A. B x: U i" and - Prod_intro: "\\x. x: A \ b x: B x; A: U i\ \ \<^bold>\x. b x: \x:A. B x" and + Prod_intro: "\A: U i; \x. x: A \ b x: B x\ \ ,\\x: A. b x: TT x: A. B x" and - Prod_elim: "\f: \x:A. B x; a: A\ \ f`a: B a" and + Prod_elim: "\f: TT x: A. B x; a: A\ \ f`a: B a" and - Prod_comp: "\a: A; \x. x: A \ b x: B x\ \ (\<^bold>\x. b x)`a \ b a" and + Prod_cmp: "\\x. x: A \ b x: B x; a: A\ \ (,\\x: A. b x)`a \ b a" and - Prod_uniq: "f: \x:A. B x \ \<^bold>\x. f`x \ f" and + Prod_uniq: "f: TT x: A. B x \ ,\\x: A. f`x \ f" and \ \Congruence rules\ - Prod_form_eq: "\A: U i; B: A \ U i; C: A \ U i; \x. x: A \ B x \ C x\ \ \x:A. B x \ \x:A. C x" and + Prod_form_eq: "\A: U i; B: A \ U i; C: A \ U i; \x. x: A \ B x \ C x\ + \ TT x: A. B x \ TT x: A. C x" and - Prod_intro_eq: "\\x. x: A \ b x \ c x; A: U i\ \ \<^bold>\x. b x \ \<^bold>\x. c x" + Prod_intro_eq: "\\x. x: A \ b x \ c x; A: U i\ \ ,\\x: A. b x \ ,\\x: A. c x" text \ The Pure rules for \\\ only let us judge strict syntactic equality of object lambda expressions. -The actual definitional equality rule is @{thm Prod_intro_eq}. +The actual definitional equality rule in the type theory is @{thm Prod_intro_eq}. Note that this is a separate rule from function extensionality. - -Note that the bold lambda symbol \\<^bold>\\ used for dependent functions clashes with the proof term syntax (cf. \
2.5.2 of the Isabelle/Isar Implementation). \ lemmas Prod_form [form] lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim -lemmas Prod_comps [comp] = Prod_comp Prod_uniq - +lemmas Prod_comp [comp] = Prod_cmp Prod_uniq section \Additional definitions\ -definition compose :: "[t, t] \ t" (infixr "o" 110) where "g o f \ \<^bold>\x. g`(f`x)" - -syntax "_compose" :: "[t, t] \ t" (infixr "\" 110) -translations "g \ f" \ "g o f" +definition compose :: "[t, t, t] \ t" +where "compose A g f \ ,\\x: A. g`(f`x)" declare compose_def [comp] +syntax "_compose" :: "[t, t] \ t" (infixr "o" 110) + +parse_translation \ +let fun compose_tr ctxt tms = + let + val f::g::_ = tms; + fun domain f = @{term "A :: t"} + in + @{const compose} $ (domain f) $ f $ g + end +in + [("_compose", compose_tr)] +end +\ + +ML_val \ +Proof_Context.get_thms @{context} "compose_def" +\ + lemma compose_assoc: - assumes "A: U i" and "f: A \ B" "g: B \ C" "h: \x:C. D x" - shows "(h \ g) \ f \ h \ (g \ f)" -by (derive lems: assms Prod_intro_eq) + assumes "A: U i" "f: A \ B" "g: B \ C" "h: TT x: C. D x" + shows "(h o g) o f \ h o (g o f)" +ML_prf \ +@{thms assms}; +\ +(* (derive lems: assms Prod_intro_eq) *) lemma compose_comp: assumes "A: U i" and "\x. x: A \ b x: B" and "\x. x: B \ c x: C x" shows "(\<^bold>\x. c x) \ (\<^bold>\x. b x) \ \<^bold>\x. c (b x)" by (derive lems: assms Prod_intro_eq) -abbreviation id :: t where "id \ \<^bold>\x. x" \ \Polymorphic identity function\ +abbreviation id :: "t \ t" where "id A \ ,\\x: A. x" end -- cgit v1.2.3