diff -r bd76f0398aa9 -r 653460d3e849 quotient_def.ML --- a/quotient_def.ML Thu Nov 05 16:43:57 2009 +0100 +++ b/quotient_def.ML Fri Nov 06 09:48:37 2009 +0100 @@ -3,7 +3,7 @@ sig datatype flag = absF | repF val get_fun: flag -> (typ * typ) list -> Proof.context -> typ -> term - val make_def: binding -> term -> typ -> mixfix -> Attrib.binding -> (typ * typ) list -> + val make_def: binding -> typ -> mixfix -> Attrib.binding -> term -> Proof.context -> (term * thm) * local_theory val quotdef: (binding * typ * mixfix) * (Attrib.binding * term) -> @@ -72,23 +72,6 @@ | _ => error ("no type variables allowed")) end -fun make_def nconst_bname rhs qty mx attr qenv lthy = -let - val (arg_tys, res_ty) = strip_type qty - - val rep_fns = map (get_fun repF qenv lthy) arg_tys - val abs_fn = (get_fun absF qenv lthy) res_ty - - fun mk_fun_map t s = - Const (@{const_name "fun_map"}, dummyT) $ t $ s - - val absrep_trm = (fold_rev mk_fun_map rep_fns abs_fn $ rhs) - |> Syntax.check_term lthy -in - define nconst_bname mx attr absrep_trm lthy -end - - (* returns all subterms where two types differ *) fun diff (T, S) Ds = case (T, S) of @@ -104,19 +87,19 @@ (* sanity check that the calculated quotient environment matches with the stored quotient environment. *) -fun error_msg lthy (qty, rty) = -let - val qtystr = quote (Syntax.string_of_typ lthy qty) - val rtystr = quote (Syntax.string_of_typ lthy rty) -in - error (implode ["Quotient type ", qtystr, " does not match with ", rtystr]) -end - -fun sanity_chk lthy qenv = +fun sanity_chk qenv lthy = let val global_qenv = Quotient_Info.mk_qenv lthy val thy = ProofContext.theory_of lthy + fun error_msg lthy (qty, rty) = + let + val qtystr = quote (Syntax.string_of_typ lthy qty) + val rtystr = quote (Syntax.string_of_typ lthy rty) + in + error (implode ["Quotient type ", qtystr, " does not match with ", rtystr]) + end + fun is_inst (qty, rty) (qty', rty') = if Sign.typ_instance thy (qty, qty') then let @@ -134,17 +117,43 @@ map chk_inst qenv end +fun make_def nconst_bname qty mx attr rhs lthy = +let + val (arg_tys, res_ty) = strip_type qty + + val rty = fastype_of rhs + val qenv = distinct (op=) (diff (qty, rty) []) + + val _ = sanity_chk qenv lthy + + val rep_fns = map (get_fun repF qenv lthy) arg_tys + val abs_fn = (get_fun absF qenv lthy) res_ty + + fun mk_fun_map t s = + Const (@{const_name "fun_map"}, dummyT) $ t $ s + + val absrep_trm = (fold_rev mk_fun_map rep_fns abs_fn $ rhs) + |> Syntax.check_term lthy +in + define nconst_bname mx attr absrep_trm lthy +end + +(* interface and syntax setup *) + +(* the ML-interface takes a 5-tuple consisting of *) +(* *) +(* - the name of the constant to be lifted *) +(* - its type *) +(* - its mixfix annotation *) +(* - a meta-equation defining the constant, *) +(* and the attributes of for this meta-equality *) fun quotdef ((bind, qty, mx), (attr, prop)) lthy = let val (_, prop') = PrimitiveDefs.dest_def lthy (K true) (K false) (K false) prop val (_, rhs) = PrimitiveDefs.abs_def prop' - - val rty = fastype_of rhs - val qenv = distinct (op=) (diff (qty, rty) []) in - sanity_chk lthy qenv; - make_def bind rhs qty mx attr qenv lthy + make_def bind qty mx attr rhs lthy end fun quotdef_cmd ((bind, qtystr, mx), (attr, propstr)) lthy =