From f46df86db9308dde29e0e5f97f54546ea1dc34bf Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Mon, 18 Jan 2021 23:49:13 +0000 Subject: Swapped notation for metas (now ?) and holes (now {}), other notation and name changes. --- spartan/core/Spartan.thy | 85 ++++++++++++++++++++++-------------------------- 1 file changed, 38 insertions(+), 47 deletions(-) (limited to 'spartan/core/Spartan.thy') diff --git a/spartan/core/Spartan.thy b/spartan/core/Spartan.thy index 6b2ed58..ffe3778 100644 --- a/spartan/core/Spartan.thy +++ b/spartan/core/Spartan.thy @@ -9,23 +9,18 @@ keywords "Theorem" "Lemma" "Corollary" "Proposition" "Definition" :: thy_goal_stmt and "assuming" :: prf_asm % "proof" and "focus" "\<^item>" "\<^enum>" "\" "\" "~" :: prf_script_goal % "proof" and - "congruence" "print_coercions" :: thy_decl and + "calc" "print_coercions" :: thy_decl and "rhs" "def" "vars" :: quasi_command begin - -section \Prelude\ - -paragraph \Set up the meta-logic just so.\ - -paragraph \Notation.\ +section \Notation\ declare [[eta_contract=false]] text \ - Rebind notation for meta-lambdas since we want to use `\` for the object - lambdas. Meta-functions now use the binder `fn`. +Rebind notation for meta-lambdas since we want to use \\\ for the object +lambdas. Metafunctions now use the binder \fn\. \ setup \ let @@ -41,34 +36,31 @@ in end \ -syntax "_dollar" :: \logic \ logic \ logic\ (infixr "$" 3) +syntax "_app" :: \logic \ logic \ logic\ (infixr "$" 3) translations "a $ b" \ "a (b)" abbreviation (input) K where "K x \ fn _. x" -paragraph \ - Deeply embed dependent type theory: meta-type of expressions, and typing - judgment. + +section \Metalogic\ + +text \ +HOAS embedding of dependent type theory: metatype of expressions, and typing +judgment. \ typedecl o consts has_type :: \o \ o \ prop\ ("(2_:/ _)" 999) -text \Type annotations for type-checking and inference.\ - -definition anno :: \o \ o \ o\ ("'(_ : _')") - where "(a : A) \ a" if "a: A" - -lemma anno: "a: A \ (a : A): A" - by (unfold anno_def) - section \Axioms\ subsection \Universes\ -typedecl lvl \ \Universe levels\ +text \\-many cumulative Russell universes.\ + +typedecl lvl axiomatization O :: \lvl\ and @@ -81,16 +73,15 @@ axiomatization axiomatization U :: \lvl \ o\ where Ui_in_Uj: "i < j \ U i: U j" and - in_Uj_if_in_Ui: "A: U i \ i < j \ A: U j" + U_cumul: "A: U i \ i < j \ A: U j" lemma Ui_in_USi: "U i: U (S i)" by (rule Ui_in_Uj, rule lt_S) -lemma in_USi_if_in_Ui: +lemma U_lift: "A: U i \ A: U (S i)" - by (erule in_Uj_if_in_Ui, rule lt_S) - + by (erule U_cumul, rule lt_S) subsection \\-type\ @@ -134,7 +125,6 @@ axiomatization where lam_cong: "\\x. x: A \ b x \ c x; A: U i\ \ \x: A. b x \ \x: A. c x" - subsection \\-type\ axiomatization @@ -207,7 +197,7 @@ lemma sub: shows "a: A'" using assms by simp -\ \Basic substitution of definitional equalities\ +\ \Basic rewriting of computational equality\ ML_file \~~/src/Tools/misc_legacy.ML\ ML_file \~~/src/Tools/IsaPlanner/isand.ML\ ML_file \~~/src/Tools/IsaPlanner/rw_inst.ML\ @@ -305,7 +295,7 @@ method_setup this = subsection \Rewriting\ -consts rewrite_HOLE :: "'a::{}" ("\") +consts rewrite_hole :: "'a::{}" ("\") lemma eta_expand: fixes f :: "'a::{} \ 'b::{}" @@ -331,28 +321,28 @@ lemma imp_cong_eq: ML_file \~~/src/HOL/Library/cconv.ML\ ML_file \rewrite.ML\ -\ \\reduce\ computes terms via judgmental equalities\ +\ \\reduce\ simplifies terms via computational equalities\ method reduce uses add = changed \repeat_new \(simp add: comp add | subst comp); typechk?\\ -subsection \Congruence relations\ +subsection \Calculational reasoning\ consts "rhs" :: \'a\ ("..") -ML_file \congruence.ML\ +ML_file \calc.ML\ section \Implicits\ text \ - \?\ is used to mark implicit arguments in definitions, while \{}\ is expanded + \{}\ is used to mark implicit arguments in definitions, while \?\ is expanded immediately for elaboration in statements. \ consts - iarg :: \'a\ ("?") - hole :: \'b\ ("{}") + iarg :: \'a\ ("{}") + hole :: \'b\ ("?") ML_file \implicits.ML\ @@ -363,11 +353,12 @@ ML \val _ = Context.>> (Syntax_Phases.term_check 1 "" Implicits.make_holes text \Automatically insert inhabitation judgments where needed:\ syntax inhabited :: \o \ prop\ ("(_)") -translations "inhabited A" \ "CONST has_type {} A" +translations "inhabited A" \ "CONST has_type ? A" + -text \Implicit lambdas\ +subsection \Implicit lambdas\ -definition lam_i where [implicit]: "lam_i f \ lam ? f" +definition lam_i where [implicit]: "lam_i f \ lam {} f" syntax "_lam_i" :: \idts \ o \ o\ ("(2\_./ _)" 30) @@ -410,7 +401,7 @@ Lemma refine_codomain: Lemma lift_universe_codomain: assumes "A: U i" "f: A \ U j" shows "f: A \ U (S j)" - using in_USi_if_in_Ui[of "f {}"] + using U_lift by (rule refine_codomain) subsection \Function composition\ @@ -460,10 +451,10 @@ Lemma funcomp_apply_comp [comp]: shows "(g \\<^bsub>A\<^esub> f) x \ g (f x)" unfolding funcomp_def by reduce -text \Notation:\ +subsection \Notation\ definition funcomp_i (infixr "\" 120) - where [implicit]: "funcomp_i g f \ g \\<^bsub>?\<^esub> f" + where [implicit]: "funcomp_i g f \ g \\<^bsub>{}\<^esub> f" translations "g \ f" \ "g \\<^bsub>A\<^esub> f" @@ -520,10 +511,10 @@ Lemma snd_comp [comp]: subsection \Notation\ definition fst_i ("fst") - where [implicit]: "fst \ Spartan.fst ? ?" + where [implicit]: "fst \ Spartan.fst {} {}" definition snd_i ("snd") - where [implicit]: "snd \ Spartan.snd ? ?" + where [implicit]: "snd \ Spartan.snd {} {}" translations "fst" \ "CONST Spartan.fst A B" @@ -550,10 +541,10 @@ method snd for p::o = rule snd[where ?p=p] text \Double projections:\ -definition [implicit]: "p\<^sub>1\<^sub>1 p \ Spartan.fst ? ? (Spartan.fst ? ? p)" -definition [implicit]: "p\<^sub>1\<^sub>2 p \ Spartan.snd ? ? (Spartan.fst ? ? p)" -definition [implicit]: "p\<^sub>2\<^sub>1 p \ Spartan.fst ? ? (Spartan.snd ? ? p)" -definition [implicit]: "p\<^sub>2\<^sub>2 p \ Spartan.snd ? ? (Spartan.snd ? ? p)" +definition [implicit]: "p\<^sub>1\<^sub>1 p \ Spartan.fst {} {} (Spartan.fst {} {} p)" +definition [implicit]: "p\<^sub>1\<^sub>2 p \ Spartan.snd {} {} (Spartan.fst {} {} p)" +definition [implicit]: "p\<^sub>2\<^sub>1 p \ Spartan.fst {} {} (Spartan.snd {} {} p)" +definition [implicit]: "p\<^sub>2\<^sub>2 p \ Spartan.snd {} {} (Spartan.snd {} {} p)" translations "CONST p\<^sub>1\<^sub>1 p" \ "fst (fst p)" -- cgit v1.2.3