wenzelm@18620: (* Title: Pure/Isar/specification.ML wenzelm@18620: ID: $Id$ wenzelm@18620: Author: Makarius wenzelm@18620: wenzelm@18810: Common theory/locale specifications --- with type-inference and wenzelm@18810: toplevel polymorphism. wenzelm@18620: *) wenzelm@18620: wenzelm@18620: signature SPECIFICATION = wenzelm@18620: sig wenzelm@18771: val read_specification: (string * string option * mixfix) list -> wenzelm@18954: ((string * Attrib.src list) * string list) list -> local_theory -> wenzelm@18771: (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) * wenzelm@18954: local_theory wenzelm@18771: val cert_specification: (string * typ option * mixfix) list -> wenzelm@18954: ((string * Attrib.src list) * term list) list -> local_theory -> wenzelm@18771: (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) * wenzelm@18954: local_theory wenzelm@18954: val axiomatization: (string * string option * mixfix) list -> wenzelm@18954: ((bstring * Attrib.src list) * string list) list -> local_theory -> wenzelm@18954: (term list * (bstring * thm list) list) * local_theory wenzelm@18954: val axiomatization_i: (string * typ option * mixfix) list -> wenzelm@18954: ((bstring * Attrib.src list) * term list) list -> local_theory -> wenzelm@18954: (term list * (bstring * thm list) list) * local_theory wenzelm@18954: val definition: wenzelm@18786: ((string * string option * mixfix) option * ((string * Attrib.src list) * string)) list -> wenzelm@18954: local_theory -> (term * (bstring * thm)) list * local_theory wenzelm@18954: val definition_i: wenzelm@18786: ((string * typ option * mixfix) option * ((string * Attrib.src list) * term)) list -> wenzelm@18954: local_theory -> (term * (bstring * thm)) list * local_theory wenzelm@19080: val abbreviation: bool -> ((string * string option * mixfix) option * string) list -> wenzelm@19080: local_theory -> local_theory wenzelm@19080: val abbreviation_i: bool -> ((string * typ option * mixfix) option * term) list -> wenzelm@19080: local_theory -> local_theory wenzelm@18620: end; wenzelm@18620: wenzelm@18620: structure Specification: SPECIFICATION = wenzelm@18620: struct wenzelm@18620: wenzelm@18620: (* prepare specification *) wenzelm@18620: wenzelm@18828: fun prep_specification prep_vars prep_propp prep_att raw_vars raw_specs ctxt = wenzelm@18620: let wenzelm@18670: val thy = ProofContext.theory_of ctxt; wenzelm@18620: wenzelm@18670: val (vars, vars_ctxt) = ctxt |> prep_vars raw_vars; wenzelm@18670: val (xs, params_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars; wenzelm@18670: val ((specs, vs), specs_ctxt) = wenzelm@18620: prep_propp (params_ctxt, map (map (rpair ([], [])) o snd) raw_specs) wenzelm@18620: |> swap |>> map (map fst) wenzelm@18771: ||>> fold_map ProofContext.inferred_param xs; wenzelm@18620: wenzelm@18771: val params = vs ~~ map #3 vars; wenzelm@18620: val names = map (fst o fst) raw_specs; wenzelm@18670: val atts = map (map (prep_att thy) o snd o fst) raw_specs; wenzelm@18620: in ((params, (names ~~ atts) ~~ specs), specs_ctxt) end; wenzelm@18620: wenzelm@18620: fun read_specification x = wenzelm@18771: prep_specification ProofContext.read_vars ProofContext.read_propp Attrib.intern_src x; wenzelm@18620: fun cert_specification x = wenzelm@18670: prep_specification ProofContext.cert_vars ProofContext.cert_propp (K I) x; wenzelm@18620: wenzelm@18620: wenzelm@18771: (* axiomatization *) wenzelm@18620: wenzelm@18954: fun gen_axioms prep raw_vars raw_specs ctxt = wenzelm@18620: let wenzelm@18786: val (vars, specs) = fst (prep raw_vars raw_specs ctxt); wenzelm@18828: val cs = map fst vars; wenzelm@18880: val spec_frees = member (op =) (fold (fold Term.add_frees o snd) specs []); wenzelm@18786: wenzelm@18880: val (consts, consts_ctxt) = ctxt |> LocalTheory.consts_restricted spec_frees vars; wenzelm@18828: val subst = Term.subst_atomic (map Free cs ~~ consts); wenzelm@18786: wenzelm@18771: val (axioms, axioms_ctxt) = wenzelm@18771: consts_ctxt wenzelm@18771: |> LocalTheory.axioms (specs |> map (fn (a, props) => (a, map subst props))) wenzelm@18786: ||> LocalTheory.theory (Theory.add_finals_i false (map Term.head_of consts)); wenzelm@18954: val _ = LocalTheory.print_consts ctxt spec_frees cs; wenzelm@18954: in ((consts, axioms), axioms_ctxt) end; wenzelm@18786: wenzelm@18954: val axiomatization = gen_axioms read_specification; wenzelm@18954: val axiomatization_i = gen_axioms cert_specification; wenzelm@18620: wenzelm@18786: wenzelm@18786: (* definition *) wenzelm@18786: wenzelm@18954: fun gen_defs prep args ctxt = wenzelm@18786: let wenzelm@18954: fun define (raw_var, (raw_a, raw_prop)) ctxt' = wenzelm@18786: let wenzelm@18954: val (vars, [(a, [prop])]) = fst (prep (the_list raw_var) [(raw_a, [raw_prop])] ctxt'); wenzelm@18954: val (((x, T), rhs), prove) = LocalDefs.derived_def ctxt' true prop; wenzelm@18786: val mx = (case vars of [] => NoSyn | [((x', _), mx)] => wenzelm@18786: if x = x' then mx wenzelm@18786: else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote x')); wenzelm@18810: in wenzelm@18954: ctxt' wenzelm@18810: |> LocalTheory.def_finish prove ((x, mx), (a, rhs)) wenzelm@18810: |>> pair (x, T) wenzelm@18810: end; wenzelm@18786: wenzelm@18828: val ((cs, defs), defs_ctxt) = ctxt |> fold_map define args |>> split_list; wenzelm@18880: val def_frees = member (op =) (fold (Term.add_frees o fst) defs []); wenzelm@18954: val _ = LocalTheory.print_consts ctxt def_frees cs; wenzelm@18954: in (defs, defs_ctxt) end; wenzelm@18786: wenzelm@18954: val definition = gen_defs read_specification; wenzelm@18954: val definition_i = gen_defs cert_specification; wenzelm@18786: wenzelm@19080: wenzelm@19080: (* abbreviation *) wenzelm@19080: wenzelm@19080: fun gen_abbrevs prep revert args ctxt = wenzelm@19080: let wenzelm@19080: fun abbrev (raw_var, raw_prop) ctxt' = wenzelm@19080: let wenzelm@19080: val (vars, [(_, [prop])]) = fst (prep (the_list raw_var) [(("", []), [raw_prop])] ctxt'); wenzelm@19080: val ((x, T), rhs) = #1 (LocalDefs.derived_def ctxt' false prop); wenzelm@19080: val mx = (case vars of [] => NoSyn | [((x', _), mx)] => wenzelm@19080: if x = x' then mx wenzelm@19080: else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote x')); wenzelm@19080: in wenzelm@19080: ctxt' wenzelm@19080: |> LocalTheory.abbrev revert ((x, mx), rhs) wenzelm@19080: |> pair (x, T) wenzelm@19080: end; wenzelm@19080: wenzelm@19080: val (cs, abbrs_ctxt) = ctxt |> fold_map abbrev args; wenzelm@19080: val _ = LocalTheory.print_consts ctxt (K false) cs; wenzelm@19080: in abbrs_ctxt end; wenzelm@19080: wenzelm@19080: val abbreviation = gen_abbrevs read_specification; wenzelm@19080: val abbreviation_i = gen_abbrevs cert_specification; wenzelm@19080: wenzelm@18620: end;