# HG changeset patch # User wenzelm # Date 1267962422 -3600 # Node ID 06197484c6ad1975bac1ab4115d8cfa959c0acf4 # Parent 9c818cab0dd0d855c0db78413ffc53ca4ca17735 separate structure Typedecl; diff -r 9c818cab0dd0 -r 06197484c6ad src/HOL/Boogie/Tools/boogie_loader.ML --- a/src/HOL/Boogie/Tools/boogie_loader.ML Sun Mar 07 12:19:47 2010 +0100 +++ b/src/HOL/Boogie/Tools/boogie_loader.ML Sun Mar 07 12:47:02 2010 +0100 @@ -88,7 +88,7 @@ let val args = Name.variant_list [] (replicate arity "'") val (T, thy') = - Object_Logic.typedecl (Binding.name isa_name, args, mk_syntax name arity) thy + Typedecl.typedecl (Binding.name isa_name, args, mk_syntax name arity) thy val type_name = fst (Term.dest_Type T) in (((name, type_name), log_new name type_name), thy') end) end diff -r 9c818cab0dd0 -r 06197484c6ad src/HOL/Tools/typedef.ML --- a/src/HOL/Tools/typedef.ML Sun Mar 07 12:19:47 2010 +0100 +++ b/src/HOL/Tools/typedef.ML Sun Mar 07 12:47:02 2010 +0100 @@ -130,7 +130,7 @@ in Drule.export_without_context (Drule.equal_elim_rule2 OF [goal_eq, th]) end; fun typedef_result inhabited = - Object_Logic.typedecl (tname, vs, mx) + Typedecl.typedecl (tname, vs, mx) #> snd #> Sign.add_consts_i [(Rep_name, newT --> oldT, NoSyn), diff -r 9c818cab0dd0 -r 06197484c6ad src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Sun Mar 07 12:19:47 2010 +0100 +++ b/src/Pure/IsaMakefile Sun Mar 07 12:47:02 2010 +0100 @@ -72,18 +72,19 @@ Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML \ Isar/rule_insts.ML Isar/runtime.ML Isar/skip_proof.ML \ Isar/spec_parse.ML Isar/spec_rules.ML Isar/specification.ML \ - Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML \ - ML/ml_antiquote.ML ML/ml_compiler.ML ML/ml_compiler_polyml-5.3.ML \ - ML/ml_context.ML ML/ml_env.ML ML/ml_lex.ML ML/ml_parse.ML \ - ML/ml_syntax.ML ML/ml_thms.ML ML-Systems/install_pp_polyml.ML \ - ML-Systems/install_pp_polyml-5.3.ML ML-Systems/use_context.ML \ - Proof/extraction.ML Proof/proof_rewrite_rules.ML \ - Proof/proof_syntax.ML Proof/proofchecker.ML Proof/reconstruct.ML \ - ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML \ - ProofGeneral/pgip_isabelle.ML ProofGeneral/pgip_markup.ML \ - ProofGeneral/pgip_output.ML ProofGeneral/pgip_parser.ML \ - ProofGeneral/pgip_tests.ML ProofGeneral/pgip_types.ML \ - ProofGeneral/preferences.ML ProofGeneral/proof_general_emacs.ML \ + Isar/theory_target.ML Isar/toplevel.ML Isar/typedecl.ML \ + Isar/value_parse.ML ML/ml_antiquote.ML ML/ml_compiler.ML \ + ML/ml_compiler_polyml-5.3.ML ML/ml_context.ML ML/ml_env.ML \ + ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML ML/ml_thms.ML \ + ML-Systems/install_pp_polyml.ML ML-Systems/install_pp_polyml-5.3.ML \ + ML-Systems/use_context.ML Proof/extraction.ML \ + Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML \ + Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/pgip.ML \ + ProofGeneral/pgip_input.ML ProofGeneral/pgip_isabelle.ML \ + ProofGeneral/pgip_markup.ML ProofGeneral/pgip_output.ML \ + ProofGeneral/pgip_parser.ML ProofGeneral/pgip_tests.ML \ + ProofGeneral/pgip_types.ML ProofGeneral/preferences.ML \ + ProofGeneral/proof_general_emacs.ML \ ProofGeneral/proof_general_pgip.ML Pure.thy ROOT.ML Syntax/ast.ML \ Syntax/lexicon.ML Syntax/mixfix.ML Syntax/parser.ML \ Syntax/printer.ML Syntax/simple_syntax.ML Syntax/syn_ext.ML \ diff -r 9c818cab0dd0 -r 06197484c6ad src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Sun Mar 07 12:19:47 2010 +0100 +++ b/src/Pure/Isar/isar_syn.ML Sun Mar 07 12:47:02 2010 +0100 @@ -105,7 +105,7 @@ val _ = OuterSyntax.command "typedecl" "type declaration" K.thy_decl (P.type_args -- P.binding -- P.opt_mixfix >> (fn ((args, a), mx) => - Toplevel.theory (Object_Logic.typedecl (a, args, mx) #> snd))); + Toplevel.theory (Typedecl.typedecl (a, args, mx) #> snd))); val _ = OuterSyntax.command "types" "declare type abbreviations" K.thy_decl diff -r 9c818cab0dd0 -r 06197484c6ad src/Pure/Isar/object_logic.ML --- a/src/Pure/Isar/object_logic.ML Sun Mar 07 12:19:47 2010 +0100 +++ b/src/Pure/Isar/object_logic.ML Sun Mar 07 12:47:02 2010 +0100 @@ -8,7 +8,6 @@ sig val get_base_sort: theory -> sort option val add_base_sort: sort -> theory -> theory - val typedecl: binding * string list * mixfix -> theory -> typ * theory val add_judgment: binding * typ * mixfix -> theory -> theory val add_judgment_cmd: binding * string * mixfix -> theory -> theory val judgment_name: theory -> string @@ -82,24 +81,6 @@ else (SOME S, judgment, atomize_rulify)); -(* typedecl *) - -fun typedecl (b, vs, mx) thy = - let - val base_sort = get_base_sort thy; - val _ = has_duplicates (op =) vs andalso - error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b)); - val name = Sign.full_name thy b; - val n = length vs; - val T = Type (name, map (fn v => TFree (v, [])) vs); - in - thy - |> Sign.add_types [(b, n, mx)] - |> (case base_sort of NONE => I | SOME S => AxClass.axiomatize_arity (name, replicate n S, S)) - |> pair T - end; - - (* add judgment *) local diff -r 9c818cab0dd0 -r 06197484c6ad src/Pure/Isar/typedecl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Isar/typedecl.ML Sun Mar 07 12:47:02 2010 +0100 @@ -0,0 +1,31 @@ +(* Title: Pure/Isar/typedecl.ML + Author: Makarius + +Type declarations (within the object logic). +*) + +signature TYPEDECL = +sig + val typedecl: binding * string list * mixfix -> theory -> typ * theory +end; + +structure Typedecl: TYPEDECL = +struct + +fun typedecl (b, vs, mx) thy = + let + val base_sort = Object_Logic.get_base_sort thy; + val _ = has_duplicates (op =) vs andalso + error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b)); + val name = Sign.full_name thy b; + val n = length vs; + val T = Type (name, map (fn v => TFree (v, [])) vs); + in + thy + |> Sign.add_types [(b, n, mx)] + |> (case base_sort of NONE => I | SOME S => AxClass.axiomatize_arity (name, replicate n S, S)) + |> pair T + end; + +end; + diff -r 9c818cab0dd0 -r 06197484c6ad src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Sun Mar 07 12:19:47 2010 +0100 +++ b/src/Pure/ROOT.ML Sun Mar 07 12:47:02 2010 +0100 @@ -218,6 +218,7 @@ use "Isar/spec_parse.ML"; use "Isar/spec_rules.ML"; use "Isar/specification.ML"; +use "Isar/typedecl.ML"; use "Isar/constdefs.ML"; (*toplevel transactions*)