From afef9e63cb11267dc69b714a8b76415d75e2dd37 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Sun, 10 Feb 2019 02:59:30 +0100 Subject: Type information storage functionality (assume_type, assume_types keywords) done! Inference and pretty-printing for function composition done. --- Prod.thy | 58 ++++++++++++++++++++++++++++++++++++---------------------- 1 file changed, 36 insertions(+), 22 deletions(-) (limited to 'Prod.thy') diff --git a/Prod.thy b/Prod.thy index 0de7d89..ce785d6 100644 --- a/Prod.thy +++ b/Prod.thy @@ -1,6 +1,6 @@ (******** Isabelle/HoTT: Dependent product (dependent function) -Jan 2019 +Feb 2019 ********) @@ -20,13 +20,13 @@ axiomatization syntax "_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) + "_lam" :: "[idt, t, t] \ t" ("(3\'(_: _')./ _)" 30) + "_lam'" :: "[idt, t, t] \ t" ("(3\_: _./ _)" 30) translations - "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)" + "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}. @@ -40,22 +40,22 @@ where "A \ B \ TT(_: A). B" axiomatization where \ \Type rules\ - Prod_form: "\A: U i; \x. x: A \ B x: U i\ \ TT x: A. B x: U i" and + Prod_form: "\A: U i; B: A \ U i\ \ TT x: A. B x: U i" and - Prod_intro: "\A: U i; \x. x: A \ b x: B x\ \ ,\\x: A. b x: TT 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: TT x: A. B x; a: A\ \ f`a: B a" and - Prod_cmp: "\\x. x: A \ b x: B x; a: A\ \ (,\\x: A. 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: TT x: A. B x \ ,\\x: A. 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\ \ TT x: A. B x \ TT x: A. C x" and - Prod_intro_eq: "\\x. x: A \ b x \ c x; A: U i\ \ ,\\x: A. b x \ ,\\x: A. 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. @@ -71,23 +71,23 @@ lemmas Prod_cong [cong] = Prod_form_eq Prod_intro_eq section \Function composition\ definition compose :: "[t, t, t] \ t" -where "compose A g f \ ,\\x: A. g`(f`x)" +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 fun compose_tr ctxt [g, f] = let - val g :: f :: _ = tms |> map (Typing.tm_of_ptm ctxt) + val [g, f] = [g, f] |> map (Typing.prep_term @{context}) val dom = case f of Const ("Prod.lam", _) $ T $ _ => T - | _ => (case Typing.get_typing f (Typing.typing_assms ctxt) of + | Const ("Prod.compose", _) $ T $ _ $ _ => T + | _ => (case Typing.get_typing ctxt f of SOME (Const ("Prod.Prod", _) $ T $ _) => T - | SOME _ => Exn.error "Can't compose with a non-function" - | NONE => Exn.error "Cannot infer domain of composition: please state this explicitly") + | SOME t => (@{print} t; Exn.error "Can't compose with a non-function") + | NONE => Exn.error "Cannot infer domain of composition; please state this explicitly") in @{const compose} $ dom $ g $ f end @@ -96,16 +96,30 @@ in end \ +text \Pretty-printing switch for composition; hides domain type information.\ + +ML \val pretty_compose = Attrib.setup_config_bool @{binding "pretty_compose"} (K true)\ + +print_translation \ +let fun compose_tr' ctxt [A, g, f] = + if Config.get ctxt pretty_compose + then Syntax.const @{syntax_const "_compose"} $ g $ f + else Const ("compose", Syntax.read_typ ctxt "t \ t \ t") $ A $ g $ f +in + [(@{const_syntax compose}, compose_tr')] +end +\ + lemma compose_assoc: - assumes "A: U i" "f: A \ B" "g: B \ C" "h: TT x: C. D x" + assumes "A: U i" and "f: A \ B" and "g: B \ C" and "h: TT x: C. D x" shows "compose A (compose B h g) f \ compose A h (compose A g f)" by (derive lems: assms cong) lemma compose_comp: assumes "A: U i" and "\x. x: A \ b x: B" and "\x. x: B \ c x: C x" - shows "(,\\x: B. c x) o (,\\x: A. b x) \ ,\\x: A. c (b x)" + shows "(\x: B. c x) o (\x: A. b x) \ \x: A. c (b x)" by (derive lems: assms cong) -abbreviation id :: "t \ t" where "id A \ ,\\x: A. x" +abbreviation id :: "t \ t" where "id A \ \x: A. x" end -- cgit v1.2.3