src/HOL/HOLCF/Tools/Domain/domain.ML
author huffman
Sat, 27 Nov 2010 16:08:10 -0800
changeset 40774 0437dbc127b3
parent 40737 src/HOLCF/Tools/Domain/domain.ML@2037021f034f
child 40832 4352ca878c41
permissions -rw-r--r--
moved directory src/HOLCF to src/HOL/HOLCF; added HOLCF theories to src/HOL/IsaMakefile;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
40040
3adb92ee2f22 rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
huffman
parents: 40026
diff changeset
     1
(*  Title:      HOLCF/Tools/Domain/domain.ML
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     David von Oheimb
35794
8cd7134275cc use headers consistently
huffman
parents: 35776
diff changeset
     3
    Author:     Brian Huffman
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
Theory extender for domain command, including theory syntax.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
40040
3adb92ee2f22 rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
huffman
parents: 40026
diff changeset
     8
signature DOMAIN =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
sig
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    10
  val add_domain_cmd:
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    11
      ((string * string option) list * binding * mixfix *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    12
       (binding * (bool * binding option * string) list * mixfix) list) list
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    13
      -> theory -> theory
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    14
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    15
  val add_domain:
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    16
      ((string * sort) list * binding * mixfix *
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    17
       (binding * (bool * binding option * typ) list * mixfix) list) list
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    18
      -> theory -> theory
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    19
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    20
  val add_new_domain_cmd:
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    21
      ((string * string option) list * binding * mixfix *
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    22
       (binding * (bool * binding option * string) list * mixfix) list) list
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    23
      -> theory -> theory
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    24
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    25
  val add_new_domain:
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    26
      ((string * sort) list * binding * mixfix *
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    27
       (binding * (bool * binding option * typ) list * mixfix) list) list
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    28
      -> theory -> theory
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
40040
3adb92ee2f22 rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
huffman
parents: 40026
diff changeset
    31
structure Domain :> DOMAIN =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
40026
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    34
open HOLCF_Library;
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    35
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    36
fun first  (x,_,_) = x;
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    37
fun second (_,x,_) = x;
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    38
fun third  (_,_,x) = x;
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
    39
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
(* ----- calls for building new thy and thms -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
    42
type info =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
    43
     Domain_Take_Proofs.iso_info list * Domain_Take_Proofs.take_induct_info;
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
    44
40485
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    45
fun add_arity ((b, sorts, mx), sort) thy : theory =
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    46
  thy
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    47
  |> Sign.add_types [(b, length sorts, mx)]
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    48
  |> AxClass.axiomatize_arity (Sign.full_name thy b, sorts, sort);
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    49
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30916
diff changeset
    50
fun gen_add_domain
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    51
    (prep_sort : theory -> 'a -> sort)
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    52
    (prep_typ : theory -> (string * sort) list -> 'b -> typ)
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
    53
    (add_isos : (binding * mixfix * (typ * typ)) list -> theory -> info * theory)
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
    54
    (arg_sort : bool -> sort)
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    55
    (raw_specs : ((string * 'a) list * binding * mixfix *
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    56
               (binding * (bool * binding option * 'b) list * mixfix) list) list)
35516
59550ec4921d get rid of primes on thy variables
huffman
parents: 35498
diff changeset
    57
    (thy : theory) =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    58
  let
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
    59
    val dtnvs : (binding * typ list * mixfix) list =
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
    60
      let
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    61
        fun prep_tvar (a, s) = TFree (a, prep_sort thy s);
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
    62
      in
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    63
        map (fn (vs, dbind, mx, _) =>
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    64
                (dbind, map prep_tvar vs, mx)) raw_specs
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
    65
      end;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    66
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
    67
    fun thy_arity (dbind, tvars, mx) =
40485
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    68
      ((dbind, map (snd o dest_TFree) tvars, mx), arg_sort false);
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    69
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    70
    (* this theory is used just for parsing and error checking *)
35516
59550ec4921d get rid of primes on thy variables
huffman
parents: 35498
diff changeset
    71
    val tmp_thy = thy
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    72
      |> Theory.copy
40485
0150614948aa refactor tmp_thy code
huffman
parents: 40329
diff changeset
    73
      |> fold (add_arity o thy_arity) dtnvs;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    74
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    75
    val dbinds : binding list =
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
    76
        map (fn (_,dbind,_,_) => dbind) raw_specs;
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    77
    val raw_rhss :
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
    78
        (binding * (bool * binding option * 'b) list * mixfix) list list =
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
    79
        map (fn (_,_,_,cons) => cons) raw_specs;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    80
    val dtnvs' : (string * typ list) list =
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
    81
        map (fn (dbind, vs, mx) => (Sign.full_name thy dbind, vs)) dtnvs;
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    82
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    83
    val all_cons = map (Binding.name_of o first) (flat raw_rhss);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    84
    val test_dupl_cons =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    85
      case duplicates (op =) all_cons of 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    86
        [] => false | dups => error ("Duplicate constructors: " 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    87
                                      ^ commas_quote dups);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    88
    val all_sels =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    89
      (map Binding.name_of o map_filter second o maps second) (flat raw_rhss);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    90
    val test_dupl_sels =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    91
      case duplicates (op =) all_sels of
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    92
        [] => false | dups => error("Duplicate selectors: "^commas_quote dups);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    93
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    94
    fun test_dupl_tvars s =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    95
      case duplicates (op =) (map(fst o dest_TFree)s) of
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    96
        [] => false | dups => error("Duplicate type arguments: " 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    97
                                    ^commas_quote dups);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    98
    val test_dupl_tvars' = exists test_dupl_tvars (map snd dtnvs');
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
    99
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   100
    val sorts : (string * sort) list =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   101
      let val all_sorts = map (map dest_TFree o snd) dtnvs';
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   102
      in
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   103
        case distinct (eq_set (op =)) all_sorts of
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   104
          [sorts] => sorts
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   105
        | _ => error "Mutually recursive domains must have same type parameters"
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   106
      end;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   107
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   108
    (* a lazy argument may have an unpointed type *)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   109
    (* unless the argument has a selector function *)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   110
    fun check_pcpo (lazy, sel, T) =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   111
      let val sort = arg_sort (lazy andalso is_none sel) in
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   112
        if Sign.of_sort tmp_thy (T, sort) then ()
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   113
        else error ("Constructor argument type is not of sort " ^
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   114
                    Syntax.string_of_sort_global tmp_thy sort ^ ": " ^
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   115
                    Syntax.string_of_typ_global tmp_thy T)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   116
      end;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   117
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   118
    (* test for free type variables, illegal sort constraints on rhs,
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   119
       non-pcpo-types and invalid use of recursive type;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   120
       replace sorts in type variables on rhs *)
40737
2037021f034f remove map function names from domain package theory data
huffman
parents: 40501
diff changeset
   121
    val rec_tab = Domain_Take_Proofs.get_rec_tab thy;
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   122
    fun check_rec rec_ok (T as TFree (v,_))  =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   123
        if AList.defined (op =) sorts v then T
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   124
        else error ("Free type variable " ^ quote v ^ " on rhs.")
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   125
      | check_rec rec_ok (T as Type (s, Ts)) =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   126
        (case AList.lookup (op =) dtnvs' s of
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   127
          NONE =>
40737
2037021f034f remove map function names from domain package theory data
huffman
parents: 40501
diff changeset
   128
            let val rec_ok' = rec_ok andalso Symtab.defined rec_tab s;
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   129
            in Type (s, map (check_rec rec_ok') Ts) end
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   130
        | SOME typevars =>
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   131
          if typevars <> Ts
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   132
          then error ("Recursion of type " ^ 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   133
                      quote (Syntax.string_of_typ_global tmp_thy T) ^ 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   134
                      " with different arguments")
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   135
          else if rec_ok then T
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   136
          else error ("Illegal indirect recursion of type " ^ 
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   137
                      quote (Syntax.string_of_typ_global tmp_thy T)))
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   138
      | check_rec rec_ok (TVar _) = error "extender:check_rec";
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   139
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   140
    fun prep_arg (lazy, sel, raw_T) =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   141
      let
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   142
        val T = prep_typ tmp_thy sorts raw_T;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   143
        val _ = check_rec true T;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   144
        val _ = check_pcpo (lazy, sel, T);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   145
      in (lazy, sel, T) end;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   146
    fun prep_con (b, args, mx) = (b, map prep_arg args, mx);
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   147
    fun prep_rhs cons = map prep_con cons;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   148
    val rhss : (binding * (bool * binding option * typ) list * mixfix) list list =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   149
        map prep_rhs raw_rhss;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   150
40026
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
   151
    fun mk_arg_typ (lazy, dest_opt, T) = if lazy then mk_upT T else T;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   152
    fun mk_con_typ (bind, args, mx) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   153
        if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args);
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   154
    fun mk_rhs_typ cons = foldr1 mk_ssumT (map mk_con_typ cons);
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
   155
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
   156
    val absTs : typ list = map Type dtnvs';
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   157
    val repTs : typ list = map mk_rhs_typ rhss;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   158
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   159
    val iso_spec : (binding * mixfix * (typ * typ)) list =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   160
        map (fn ((dbind, _, mx), eq) => (dbind, mx, eq))
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
   161
          (dtnvs ~~ (absTs ~~ repTs));
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   162
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   163
    val ((iso_infos, take_info), thy) = add_isos iso_spec thy;
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   164
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 39986
diff changeset
   165
    val (constr_infos, thy) =
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   166
        thy
40042
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
   167
          |> fold_map (fn ((dbind, cons), info) =>
87c72a02eaf2 simplify check_and_sort_domain; more meaningful variable names
huffman
parents: 40040
diff changeset
   168
                Domain_Constructors.add_domain_constructors dbind cons info)
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   169
             (dbinds ~~ rhss ~~ iso_infos);
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 39986
diff changeset
   170
40026
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
   171
    val (take_rews, thy) =
40097
429cd74f795f remove legacy comp_dbind option from domain package
huffman
parents: 40044
diff changeset
   172
        Domain_Induction.comp_theorems
40026
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
   173
          dbinds take_info constr_infos thy;
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   174
  in
40026
8f8f18a88685 remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
huffman
parents: 40019
diff changeset
   175
    thy
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   176
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   178
fun define_isos (spec : (binding * mixfix * (typ * typ)) list) =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   179
  let
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   180
    fun prep (dbind, mx, (lhsT, rhsT)) =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   181
      let val (dname, vs) = dest_Type lhsT;
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   182
      in (map (fst o dest_TFree) vs, dbind, mx, rhsT, NONE) end;
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   183
  in
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   184
    Domain_Isomorphism.domain_isomorphism (map prep spec)
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   185
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   186
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   187
fun pcpo_arg lazy = if lazy then @{sort cpo} else @{sort pcpo};
40501
2ed71459136e allow unpointed lazy arguments for definitional domain package
huffman
parents: 40497
diff changeset
   188
fun rep_arg lazy = if lazy then @{sort predomain} else @{sort "domain"};
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   189
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   190
fun read_sort thy (SOME s) = Syntax.read_sort_global thy s
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   191
  | read_sort thy NONE = Sign.defaultS thy;
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   192
40044
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   193
(* Adapted from src/HOL/Tools/Datatype/datatype_data.ML *)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   194
fun read_typ thy sorts str =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   195
  let
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   196
    val ctxt = ProofContext.init_global thy
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   197
      |> fold (Variable.declare_typ o TFree) sorts;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   198
  in Syntax.read_typ ctxt str end;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   199
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   200
fun cert_typ sign sorts raw_T =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   201
  let
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   202
    val T = Type.no_tvars (Sign.certify_typ sign raw_T)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   203
      handle TYPE (msg, _, _) => error msg;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   204
    val sorts' = Term.add_tfreesT T sorts;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   205
    val _ =
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   206
      case duplicates (op =) (map fst sorts') of
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   207
        [] => ()
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   208
      | dups => error ("Inconsistent sort constraints for " ^ commas dups)
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   209
  in T end;
89381a2f8864 combine check_and_sort_domain with main function; rewrite much of the error-checking code
huffman
parents: 40043
diff changeset
   210
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   211
val add_domain =
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   212
    gen_add_domain (K I) cert_typ Domain_Axioms.add_axioms pcpo_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   213
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   214
val add_new_domain =
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   215
    gen_add_domain (K I) cert_typ define_isos rep_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   216
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   217
val add_domain_cmd =
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   218
    gen_add_domain read_sort read_typ Domain_Axioms.add_axioms pcpo_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   219
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   220
val add_new_domain_cmd =
40215
d8fd7ae0254f change types of ML commands add_domain, add_new_domain to take 'sort' instead of 'string option'
huffman
parents: 40097
diff changeset
   221
    gen_add_domain read_sort read_typ define_isos rep_arg;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   222
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   223
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   224
(** outer syntax **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   225
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   226
val _ = Keyword.keyword "lazy";
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   227
val _ = Keyword.keyword "unsafe";
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   228
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30916
diff changeset
   229
val dest_decl : (bool * binding option * string) parser =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   230
  Parse.$$$ "(" |-- Scan.optional (Parse.$$$ "lazy" >> K true) false --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   231
    (Parse.binding >> SOME) -- (Parse.$$$ "::" |-- Parse.typ)  --| Parse.$$$ ")" >> Parse.triple1
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   232
    || Parse.$$$ "(" |-- Parse.$$$ "lazy" |-- Parse.typ --| Parse.$$$ ")"
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   233
    >> (fn t => (true,NONE,t))
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   234
    || Parse.typ >> (fn t => (false,NONE,t));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   235
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   236
val cons_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   237
  Parse.binding -- Scan.repeat dest_decl -- Parse.opt_mixfix;
30916
a3d2128cac92 allow infix declarations for type constructors defined with domain package
huffman
parents: 30915
diff changeset
   238
a3d2128cac92 allow infix declarations for type constructors defined with domain package
huffman
parents: 30915
diff changeset
   239
val domain_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   240
  (Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix) --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   241
    (Parse.$$$ "=" |-- Parse.enum1 "|" cons_decl);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   242
30916
a3d2128cac92 allow infix declarations for type constructors defined with domain package
huffman
parents: 30915
diff changeset
   243
val domains_decl =
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   244
  Scan.optional (Parse.$$$ "(" |-- (Parse.$$$ "unsafe" >> K true) --| Parse.$$$ ")") false --
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   245
    Parse.and_list1 domain_decl;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   246
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   247
fun mk_domain
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   248
    (unsafe : bool,
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   249
     doms : ((((string * string option) list * binding) * mixfix) *
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   250
             ((binding * (bool * binding option * string) list) * mixfix) list) list ) =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   251
  let
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   252
    val specs : ((string * string option) list * binding * mixfix *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   253
                 (binding * (bool * binding option * string) list * mixfix) list) list =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   254
        map (fn (((vs, t), mx), cons) =>
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   255
                (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   256
  in
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   257
    if unsafe
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   258
    then add_domain_cmd specs
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   259
    else add_new_domain_cmd specs
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   260
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   261
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   262
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   263
  Outer_Syntax.command "domain" "define recursive domains (HOLCF)"
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 40215
diff changeset
   264
    Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   265
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   266
end;