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/Equivalence.thy | 23 +- hott/Identity.thy | 82 +++---- spartan/core/Spartan.thy | 2 + spartan/core/elaborated_expression.ML | 402 ++++++++++++++++++++++++++++++++++ spartan/core/elaboration.ML | 76 +++++++ spartan/core/goals.ML | 2 +- spartan/core/lib.ML | 37 +++- spartan/core/typechecking.ML | 57 +++-- 8 files changed, 602 insertions(+), 79 deletions(-) create mode 100644 spartan/core/elaborated_expression.ML create mode 100644 spartan/core/elaboration.ML diff --git a/hott/Equivalence.thy b/hott/Equivalence.thy index d976677..d844b59 100644 --- a/hott/Equivalence.thy +++ b/hott/Equivalence.thy @@ -33,13 +33,16 @@ Lemma (derive) hsym: "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 intro - apply (rule pathinv) - \ by (elim H) - \ by typechk - done + "H: f ~ g" + shows "g ~ f" +unfolding homotopy_def +proof intro + fix x assume "x: A" then have "H x: f x = g x" + using \H:_\[unfolded homotopy_def] + \ \this should become unnecessary when definitional unfolding is implemented\ + by typechk + thus "g x = f x" by (rule pathinv) fact +qed typechk Lemma (derive) htrans: assumes @@ -71,9 +74,9 @@ Lemma (derive) commute_homotopy: assumes "A: U i" "B: U i" "x: A" "y: A" - "p: x =\<^bsub>A\<^esub> y" + "p: x = y" "f: A \ B" "g: A \ B" - "H: homotopy A (fn _. B) f g" + "H: f ~ g" shows "(H x) \ g[p] = f[p] \ (H y)" \ \Need this assumption unfolded for typechecking\ supply assms(8)[unfolded homotopy_def] @@ -94,7 +97,7 @@ Corollary (derive) commute_homotopy': "A: U i" "x: A" "f: A \ A" - "H: homotopy A (fn _. A) f (id A)" + "H: f ~ (id A)" shows "H (f x) = f [H x]" oops 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) diff --git a/spartan/core/Spartan.thy b/spartan/core/Spartan.thy index 2c7216e..11bdc2b 100644 --- a/spartan/core/Spartan.thy +++ b/spartan/core/Spartan.thy @@ -215,6 +215,8 @@ method_setup known = subsection \Statement commands\ ML_file \focus.ML\ +ML_file \elaboration.ML\ +ML_file \elaborated_expression.ML\ ML_file \goals.ML\ subsection \Proof methods\ diff --git a/spartan/core/elaborated_expression.ML b/spartan/core/elaborated_expression.ML new file mode 100644 index 0000000..49b7758 --- /dev/null +++ b/spartan/core/elaborated_expression.ML @@ -0,0 +1,402 @@ +(* Title: elaborated_expression.ML + Author: Joshua Chen + +A modification of parts of ~~/Pure/Isar/expression.ML to incorporate elaboration +into the assumptions mechanism. + +Most of this file is copied verbatim from the original; the only changes are the +addition of `elaborate` and a modification to `activate_i`. +*) + +structure Elaborated_Expression = struct + +local + +open Element + +fun mk_type T = (Logic.mk_type T, []) +fun mk_term t = (t, []) +fun mk_propp (p, pats) = (Type.constraint propT p, pats) + +fun dest_type (T, []) = Logic.dest_type T +fun dest_term (t, []) = t +fun dest_propp (p, pats) = (p, pats) + +fun extract_inst (_, (_, ts)) = map mk_term ts +fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)) + +fun extract_eqns es = map (mk_term o snd) es +fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs + +fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes + | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts + | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms + | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs + | extract_elem (Notes _) = [] + | extract_elem (Lazy_Notes _) = [] + +fun restore_elem (Fixes fixes, css) = + (fixes ~~ css) |> map (fn ((x, _, mx), cs) => + (x, cs |> map dest_type |> try hd, mx)) |> Fixes + | restore_elem (Constrains csts, css) = + (csts ~~ css) |> map (fn ((x, _), cs) => + (x, cs |> map dest_type |> hd)) |> Constrains + | restore_elem (Assumes asms, css) = + (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes + | restore_elem (Defines defs, css) = + (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines + | restore_elem (elem as Notes _, _) = elem + | restore_elem (elem as Lazy_Notes _, _) = elem + +fun prep (_, pats) (ctxt, t :: ts) = + let val ctxt' = Proof_Context.augment t ctxt + in + ((t, Syntax.check_props + (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), + (ctxt', ts)) + end + +fun check cs ctxt = + let + val (cs', (ctxt', _)) = fold_map prep cs + (ctxt, Syntax.check_terms + (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)) + in (cs', ctxt') end + +fun inst_morphism params ((prfx, mandatory), insts') ctxt = + let + (* parameters *) + val parm_types = map #2 params; + val type_parms = fold Term.add_tfreesT parm_types []; + + (* type inference *) + val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; + val type_parms' = fold Term.add_tvarsT parm_types' []; + val checked = + (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') + |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) + val (type_parms'', insts'') = chop (length type_parms') checked; + + (* context *) + val ctxt' = fold Proof_Context.augment checked ctxt; + val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; + val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; + + (* instantiation *) + val instT = + (type_parms ~~ map Logic.dest_type type_parms'') + |> map_filter (fn (v, T) => if TFree v = T then NONE else SOME (v, T)); + val cert_inst = + ((map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) ~~ insts'') + |> map_filter (fn (v, t) => if Free v = t then NONE else SOME (v, cert t)); + in + (Element.instantiate_normalize_morphism (map (apsnd certT) instT, cert_inst) $> + Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') + end; + +fun abs_def ctxt = + Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; + +fun declare_elem prep_var (Element.Fixes fixes) ctxt = + let val (vars, _) = fold_map prep_var fixes ctxt + in ctxt |> Proof_Context.add_fixes vars |> snd end + | declare_elem prep_var (Element.Constrains csts) ctxt = + ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd + | declare_elem _ (Element.Assumes _) ctxt = ctxt + | declare_elem _ (Element.Defines _) ctxt = ctxt + | declare_elem _ (Element.Notes _) ctxt = ctxt + | declare_elem _ (Element.Lazy_Notes _) ctxt = ctxt; + +fun parameters_of thy strict (expr, fixed) = + let + val ctxt = Proof_Context.init_global thy; + + fun reject_dups message xs = + (case duplicates (op =) xs of + [] => () + | dups => error (message ^ commas dups)); + + fun parm_eq ((p1, mx1), (p2, mx2)) = + p1 = p2 andalso + (Mixfix.equal (mx1, mx2) orelse + error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^ + Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); + + fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); + fun params_inst (loc, (prfx, (Expression.Positional insts, eqns))) = + let + val ps = params_loc loc; + val d = length ps - length insts; + val insts' = + if d < 0 then + error ("More arguments than parameters in instantiation of locale " ^ + quote (Locale.markup_name ctxt loc)) + else insts @ replicate d NONE; + val ps' = (ps ~~ insts') |> + map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); + in (ps', (loc, (prfx, (Expression.Positional insts', eqns)))) end + | params_inst (loc, (prfx, (Expression.Named insts, eqns))) = + let + val _ = + reject_dups "Duplicate instantiation of the following parameter(s): " + (map fst insts); + val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => + if AList.defined (op =) ps p then AList.delete (op =) p ps + else error (quote p ^ " not a parameter of instantiated expression")); + in (ps', (loc, (prfx, (Expression.Named insts, eqns)))) end; + fun params_expr is = + let + val (is', ps') = fold_map (fn i => fn ps => + let + val (ps', i') = params_inst i; + val ps'' = distinct parm_eq (ps @ ps'); + in (i', ps'') end) is [] + in (ps', is') end; + + val (implicit, expr') = params_expr expr; + + val implicit' = map #1 implicit; + val fixed' = map (Variable.check_name o #1) fixed; + val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; + val implicit'' = + if strict then [] + else + let + val _ = + reject_dups + "Parameter(s) declared simultaneously in expression and for clause: " + (implicit' @ fixed'); + in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; + in (expr', implicit'' @ fixed) end; + +fun parse_elem prep_typ prep_term ctxt = + Element.map_ctxt + {binding = I, + typ = prep_typ ctxt, + term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), + pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), + fact = I, + attrib = I}; + +fun prepare_stmt prep_prop prep_obtains ctxt stmt = + (case stmt of + Element.Shows raw_shows => + raw_shows |> (map o apsnd o map) (fn (t, ps) => + (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, + map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) + | Element.Obtains raw_obtains => + let + val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; + val obtains = prep_obtains thesis_ctxt thesis raw_obtains; + in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); + +fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => + let val x = Binding.name_of binding + in (binding, AList.lookup (op =) parms x, mx) end) + +fun finish_inst ctxt (loc, (prfx, inst)) = + let + val thy = Proof_Context.theory_of ctxt; + val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; + in (loc, morph) end + +fun closeup _ _ false elem = elem + | closeup (outer_ctxt, ctxt) parms true elem = + let + (* FIXME consider closing in syntactic phase -- before type checking *) + fun close_frees t = + let + val rev_frees = + Term.fold_aterms (fn Free (x, T) => + if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I + else insert (op =) (x, T) | _ => I) t []; + in fold (Logic.all o Free) rev_frees t end; + + fun no_binds [] = [] + | no_binds _ = error "Illegal term bindings in context element"; + in + (case elem of + Assumes asms => Assumes (asms |> map (fn (a, propps) => + (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) + | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => + let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) + in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) + | e => e) + end + +fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) + | finish_elem _ _ _ (Constrains _) = Constrains [] + | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) + | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) + | finish_elem _ _ _ (elem as Notes _) = elem + | finish_elem _ _ _ (elem as Lazy_Notes _) = elem + +fun check_autofix insts eqnss elems concl ctxt = + let + val inst_cs = map extract_inst insts; + val eqns_cs = map extract_eqns eqnss; + val elem_css = map extract_elem elems; + val concl_cs = (map o map) mk_propp (map snd concl); + (* Type inference *) + val (inst_cs' :: eqns_cs' :: css', ctxt') = + (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; + val (elem_css', [concl_cs']) = chop (length elem_css) css'; + in + ((map restore_inst (insts ~~ inst_cs'), + map restore_eqns (eqnss ~~ eqns_cs'), + map restore_elem (elems ~~ elem_css'), + map fst concl ~~ concl_cs'), ctxt') + end + +fun prep_full_context_statement + parse_typ parse_prop + prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr + {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt + ctxt1 + = + let + val thy = Proof_Context.theory_of ctxt1 + val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import) + fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = + let + val params = map #1 (Locale.params_of thy loc) + val inst' = prep_inst ctxt (map #1 params) inst + val parm_types' = + params |> map (#2 #> Logic.varifyT_global #> + Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> + Type_Infer.paramify_vars) + val inst'' = map2 Type.constraint parm_types' inst' + val insts' = insts @ [(loc, (prfx, inst''))] + val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt + val inst''' = insts'' |> List.last |> snd |> snd + val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt + val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 + handle ERROR msg => if null eqns then error msg else + (Locale.tracing ctxt1 + (msg ^ "\nFalling back to reading rewrites clause before activation."); + ctxt2) + val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns + val eqns' = (prep_eqns ctxt' o map snd) eqns + val eqnss' = [attrss ~~ eqns'] + val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt' + val rewrite_morph = eqns' + |> map (abs_def ctxt') + |> Variable.export_terms ctxt' ctxt + |> Element.eq_term_morphism (Proof_Context.theory_of ctxt) + |> the_default Morphism.identity + val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt + val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns'] + in (i + 1, insts', eqnss', ctxt'') end + + fun prep_elem raw_elem ctxt = + let + val ctxt' = ctxt + |> Context_Position.set_visible false + |> declare_elem prep_var_elem raw_elem + |> Context_Position.restore_visible ctxt + val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem + in (elems', ctxt') end + + val fors = fold_map prep_var_inst fixed ctxt1 |> fst + val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd + val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2) + + fun prep_stmt elems ctxt = + check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt + + val _ = + if fixed_frees then () + else + (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of + [] => () + | frees => error ("Illegal free variables in expression: " ^ + commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))) + + val ((insts, _, elems', concl), ctxt4) = ctxt3 + |> init_body + |> fold_map prep_elem raw_elems + |-> prep_stmt + + (* parameters from expression and elements *) + val xs = maps (fn Element.Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) + (Element.Fixes fors :: elems') + val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4 + val fors' = finish_fixes parms fors + val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors' + val deps = map (finish_inst ctxt5) insts + val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems' + in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end + +fun prep_inst prep_term ctxt parms (Expression.Positional insts) = + (insts ~~ parms) |> map + (fn (NONE, p) => Free (p, dummyT) + | (SOME t, _) => prep_term ctxt t) + | prep_inst prep_term ctxt parms (Expression.Named insts) = + parms |> map (fn p => + (case AList.lookup (op =) insts p of + SOME t => prep_term ctxt t | + NONE => Free (p, dummyT))) +fun parse_inst x = prep_inst Syntax.parse_term x +fun check_expr thy instances = map (apfst (Locale.check thy)) instances + +val read_full_context_statement = prep_full_context_statement + Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains + Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src + Proof_Context.read_var check_expr + +in + +fun prep_statement prep activate raw_elems raw_stmt ctxt = + let + val ((_, _, _, elems, concl), _) = + prep {strict = true, do_close = false, fixed_frees = true} + ([], []) I raw_elems raw_stmt ctxt; + val ctxt' = ctxt + |> Proof_Context.set_stmt true + |> fold_map activate elems |> #2 + |> Proof_Context.restore_stmt ctxt; + in (concl, ctxt') end + +fun elaborate ctxt assms = + let + fun subst_term env = Envir.subst_term (Envir.type_env env, Envir.term_env env) + fun elab_fact (fact, xs) assums = + let val (subst, fact') = Elab.elab_stmt ctxt assums fact in + ((fact', map (subst_term subst) xs), Thm.cterm_of ctxt fact' :: assums) + end + fun elab (b, facts) assums = + let val (facts', assums') = fold_map elab_fact facts assums + in ((b, facts'), assums') end + in + #1 (fold_map elab assms []) + end + +fun activate_i elem ctxt = + let + val elem' = + (case (map_ctxt_attrib o map) Token.init_assignable elem of + Defines defs => + Defines (defs |> map (fn ((a, atts), (t, ps)) => + ((Thm.def_binding_optional + (Binding.name (#1 (#1 (Local_Defs.cert_def ctxt (K []) t)))) a, atts), + (t, ps)))) + | Assumes assms => Assumes (elaborate ctxt assms) + | e => e); + val ctxt' = Context.proof_map (init elem') ctxt; + in ((map_ctxt_attrib o map) Token.closure elem', ctxt') end + +fun activate raw_elem ctxt = + let val elem = raw_elem |> map_ctxt + {binding = I, + typ = I, + term = I, + pattern = I, + fact = Proof_Context.get_fact ctxt, + attrib = Attrib.check_src ctxt} + in activate_i elem ctxt end; + +fun read_statement x = prep_statement read_full_context_statement activate x + +end + +end \ No newline at end of file diff --git a/spartan/core/elaboration.ML b/spartan/core/elaboration.ML new file mode 100644 index 0000000..27b6bb0 --- /dev/null +++ b/spartan/core/elaboration.ML @@ -0,0 +1,76 @@ +(* Title: elaboration.ML + Author: Joshua Chen + +Basic elaboration. +*) + +structure Elab: sig + +val elab: Proof.context -> cterm list -> term -> Envir.env +val elab_stmt: Proof.context -> cterm list -> term -> Envir.env * term + +end = struct + +(*Elaborate `tm` by solving the inference problem `tm: {}`, knowing `assums`, + which are fully elaborated, in `ctxt`. Return a substitution.*) +fun elab ctxt assums tm = + if Lib.no_vars tm + then Envir.init + else + let + val inf = Goal.init (Thm.cterm_of ctxt (Lib.typing_of_term tm)) + val res = Types.check_infer (map Thm.assume assums) 1 (ctxt, inf) + val tm' = + Thm.prop_of (#2 (Seq.hd (Seq.filter_results res))) + |> Lib.dest_prop |> Lib.term_of_typing + handle TERM ("dest_typing", [t]) => + let val typ = Logic.unprotect (Logic.strip_assums_concl t) + |> Lib.term_of_typing + in + error ("Elaboration of " ^ Syntax.string_of_term ctxt typ ^ " failed") + end + in + Seq.hd (Unify.matchers (Context.Proof ctxt) [(tm, tm')]) + end + handle Option => error + ("Elaboration of " ^ Syntax.string_of_term ctxt tm ^ " failed") + +(*Recursively elaborate a statement \x ... y. \...\ \ P x ... y by elaborating + only the types of typing judgments (in particular, does not look at judgmental + equality statements). Could also elaborate the terms of typing judgments, but + for now we assume that these are always free variables in all the cases we're + interested in.*) +fun elab_stmt ctxt assums stmt = + let + val stmt = Lib.dest_prop stmt + fun subst_term env = Envir.subst_term (Envir.type_env env, Envir.term_env env) + in + if Lib.no_vars stmt orelse Lib.is_eq stmt then + (Envir.init, stmt) + else if Lib.is_typing stmt then + let + val typ = Lib.type_of_typing stmt + val subst = elab ctxt assums typ + in (subst, subst_term subst stmt) end + else + let + fun elab' assums (x :: xs) = + let + val (env, x') = elab_stmt ctxt assums x + val assums' = + if Lib.no_vars x' then Thm.cterm_of ctxt x' :: assums else assums + in env :: elab' assums' xs end + | elab' _ [] = [] + val (prems, concl) = Lib.decompose_goal ctxt stmt + val subst = fold (curry Envir.merge) (elab' assums prems) Envir.init + val prems' = map (Thm.cterm_of ctxt o subst_term subst) prems + val subst' = + if Lib.is_typing concl then + let val typ = Lib.type_of_typing concl + in Envir.merge (subst, elab ctxt (assums @ prems') typ) end + else subst + in (subst', subst_term subst' stmt) end + end + + +end diff --git a/spartan/core/goals.ML b/spartan/core/goals.ML index 51a853e..3e2bd09 100644 --- a/spartan/core/goals.ML +++ b/spartan/core/goals.ML @@ -184,7 +184,7 @@ val schematic_theorem_cmd = gen_schematic_theorem Bundle.includes_cmd Attrib.check_src - Expression.read_statement + Elaborated_Expression.read_statement fun theorem spec descr = Outer_Syntax.local_theory_to_proof' spec ("state " ^ descr) diff --git a/spartan/core/lib.ML b/spartan/core/lib.ML index 7b93a08..392ae2e 100644 --- a/spartan/core/lib.ML +++ b/spartan/core/lib.ML @@ -6,13 +6,16 @@ val max: ('a * 'a -> bool) -> 'a list -> 'a val maxint: int list -> int (*Terms*) -val is_rigid: term -> bool val no_vars: term -> bool +val is_rigid: term -> bool +val is_eq: term -> bool +val dest_prop: term -> term val dest_eq: term -> term * term val mk_Var: string -> int -> typ -> term val lambda_var: term -> term -> term val is_typing: term -> bool +val mk_typing: term -> term -> term val dest_typing: term -> term * term val term_of_typing: term -> term val type_of_typing: term -> term @@ -21,6 +24,7 @@ val mk_Pi: term -> term -> term -> term val typing_of_term: term -> term (*Goals*) +val decompose_goal: Proof.context -> term -> term list * term val rigid_typing_concl: term -> bool (*Subterms*) @@ -49,9 +53,15 @@ val maxint = max (op >) (* Meta *) +val no_vars = not o exists_subterm is_Var + val is_rigid = not o is_Var o head_of -val no_vars = not o exists_subterm is_Var +fun is_eq (Const (\<^const_name>\Pure.eq\, _) $ _ $ _) = true + | is_eq _ = false + +fun dest_prop (Const (\<^const_name>\Pure.prop\, _) $ P) = P + | dest_prop P = P fun dest_eq (Const (\<^const_name>\Pure.eq\, _) $ t $ def) = (t, def) | dest_eq _ = error "dest_eq" @@ -72,8 +82,10 @@ fun lambda_var x tm = fun is_typing (Const (\<^const_name>\has_type\, _) $ _ $ _) = true | is_typing _ = false +fun mk_typing t T = \<^const>\has_type\ $ t $ T + fun dest_typing (Const (\<^const_name>\has_type\, _) $ t $ T) = (t, T) - | dest_typing _ = error "dest_typing" + | dest_typing t = raise TERM ("dest_typing", [t]) val term_of_typing = #1 o dest_typing val type_of_typing = #2 o dest_typing @@ -82,11 +94,28 @@ fun mk_Pi v typ body = Const (\<^const_name>\Pi\, dummyT) $ typ $ l fun typing_of_term tm = \<^const>\has_type\ $ tm $ Var (("*?", 0), \<^typ>\o\) (*The above is a bit hacky; basically we need to guarantee that the schematic - var is fresh*) + var is fresh. This works for now because no other code in the Isabelle system + or the current logic uses this identifier.*) (** Goals **) +(*Breaks a goal \x ... y. \P1; ... Pn\ \ Q into ([P1, ..., Pn], Q), fixing + \-quantified variables and keeping schematics.*) +fun decompose_goal ctxt goal = + let + val focus = + #1 (Subgoal.focus_prems ctxt 1 NONE (Thm.trivial (Thm.cterm_of ctxt goal))) + + val schematics = #2 (#schematics focus) + |> map (fn (v, ctm) => (Thm.term_of ctm, Var v)) + in + map Thm.prop_of (#prems focus) @ [Thm.term_of (#concl focus)] + |> map (subst_free schematics) + |> (fn xs => chop (length xs - 1) xs) |> apsnd the_single + end + handle List.Empty => error "Lib.decompose_goal" + fun rigid_typing_concl goal = let val concl = Logic.strip_assums_concl goal in is_typing concl andalso is_rigid (term_of_typing concl) end diff --git a/spartan/core/typechecking.ML b/spartan/core/typechecking.ML index 946ecd6..57164a1 100644 --- a/spartan/core/typechecking.ML +++ b/spartan/core/typechecking.ML @@ -11,6 +11,8 @@ val types: Proof.context -> term -> thm list val put_type: thm -> Proof.context -> Proof.context val put_types: thm list -> Proof.context -> Proof.context +val debug_typechk: bool Config.T + val known_ctac: thm list -> int -> context_tactic val check_infer: thm list -> int -> context_tactic @@ -35,6 +37,12 @@ fun put_types typings = foldr1 (op o) (map put_type typings) (* Context tactics for type-checking *) +val debug_typechk = + Attrib.setup_config_bool \<^binding>\debug_typechk\ (K false) + +fun debug_tac ctxt s = + if Config.get ctxt debug_typechk then print_tac ctxt s else all_tac + (*Solves goals without metavariables and type inference problems by resolving with facts or assumption from inline premises.*) fun known_ctac facts = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => @@ -46,42 +54,43 @@ fun known_ctac facts = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => let val ths = facts (*FIXME: Shouldn't pull nameless facts directly from context*) @ map fst (Facts.props (Proof_Context.facts_of ctxt)) - in (resolve_tac ctxt ths i ORELSE assume_tac ctxt i) st end + |> map (Simplifier.norm_hhf ctxt) + in + (debug_tac ctxt "resolve" THEN resolve_tac ctxt ths i ORELSE + debug_tac ctxt "assume" THEN assume_tac ctxt i) st + end else Seq.empty end) (*Simple bidirectional typing tactic, with some nondeterminism from backtracking search over arbitrary `typechk` rules. The current implementation does not perform any normalization.*) -local - fun check_infer_step facts i (ctxt, st) = - let - val tac = SUBGOAL (fn (goal, i) => - if Lib.rigid_typing_concl goal - then - let val net = Tactic.build_net (facts - (*MAYBE FIXME: Remove `typechk` from this list, and instead perform - definitional unfolding to (w?)hnf.*) - @(Named_Theorems.get ctxt \<^named_theorems>\typechk\) - @(Named_Theorems.get ctxt \<^named_theorems>\form\) - @(Named_Theorems.get ctxt \<^named_theorems>\intro\) - @(map #1 (Elim.rules ctxt))) - in (resolve_from_net_tac ctxt net) i end - else no_tac) - val ctxt' = ctxt - in - TACTIC_CONTEXT ctxt' (tac i st) - end -in +fun check_infer_step facts i (ctxt, st) = + let + val tac = SUBGOAL (fn (goal, i) => + if Lib.rigid_typing_concl goal + then + let val net = Tactic.build_net (facts + (*MAYBE FIXME: Remove `typechk` from this list, and instead perform + definitional unfolding to (w?)hnf.*) + @(Named_Theorems.get ctxt \<^named_theorems>\typechk\) + @(Named_Theorems.get ctxt \<^named_theorems>\form\) + @(Named_Theorems.get ctxt \<^named_theorems>\intro\) + @(map #1 (Elim.rules ctxt))) + in (resolve_from_net_tac ctxt net) i end + else no_tac) + + val ctxt' = ctxt (*TODO: Use this to store already-derived typing judgments*) + in + TACTIC_CONTEXT ctxt' (tac i st) + end fun check_infer facts i (cst as (_, st)) = - let - val ctac = known_ctac facts CORELSE' check_infer_step facts + let val ctac = known_ctac facts CORELSE' check_infer_step facts in cst |> (ctac i CTHEN CREPEAT_IN_RANGE i (Thm.nprems_of st - i) (CTRY o CREPEAT_ALL_NEW_FWD ctac)) end -end end -- cgit v1.2.3