src/HOLCF/Tools/Domain/domain_extender.ML
author huffman
Wed, 06 Oct 2010 10:49:27 -0700
changeset 39974 b525988432e9
parent 39965 da88519e6a0b
child 39986 38677db30cad
permissions -rw-r--r--
major reorganization/simplification of HOLCF type classes: removed profinite/bifinite classes and approx function; universal domain uses approx_chain locale instead of bifinite class; ideal_completion locale does not use 'take' functions, requires countable basis instead; replaced type 'udom alg_defl' with type 'sfp'; replaced class 'rep' with class 'sfp'; renamed REP('a) to SFP('a);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32126
a5042f260440 obey captialized directory names convention
haftmann
parents: 31986
diff changeset
     1
(*  Title:      HOLCF/Tools/Domain/domain_extender.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
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
signature DOMAIN_EXTENDER =
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:
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
    11
      binding ->
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    12
      ((string * string option) list * binding * mixfix *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    13
       (binding * (bool * binding option * string) list * mixfix) list) list
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    14
      -> theory -> theory
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    15
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    16
  val add_domain:
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
    17
      binding ->
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    18
      ((string * string option) list * binding * mixfix *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    19
       (binding * (bool * binding option * typ) list * mixfix) list) list
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    20
      -> theory -> theory
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    21
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    22
  val add_new_domain_cmd:
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
    23
      binding ->
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    24
      ((string * string option) list * binding * mixfix *
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    25
       (binding * (bool * binding option * string) list * mixfix) list) list
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    26
      -> theory -> theory
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    27
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    28
  val add_new_domain:
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
    29
      binding ->
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    30
      ((string * string option) list * binding * mixfix *
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    31
       (binding * (bool * binding option * typ) list * mixfix) list) list
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    32
      -> theory -> theory
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
31023
d027411c9a38 use opaque ascription for all HOLCF code
huffman
parents: 30919
diff changeset
    35
structure Domain_Extender :> DOMAIN_EXTENDER =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
open Domain_Library;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    39
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
(* ----- general testing and preprocessing of constructor list -------------- *)
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30916
diff changeset
    41
fun check_and_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
    42
    (arg_sort : bool -> sort)
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    43
    (dtnvs : (string * typ list) list)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    44
    (cons'' : (binding * (bool * binding option * typ) list * mixfix) list list)
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    45
    (thy : theory)
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31229
diff changeset
    46
    : ((string * typ list) *
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31229
diff changeset
    47
       (binding * (bool * binding option * typ) list * mixfix) list) list =
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    48
  let
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    49
    val defaultS = Sign.defaultS thy;
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    50
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    51
    val test_dupl_typs =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    52
      case duplicates (op =) (map fst dtnvs) of 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    53
        [] => false | dups => error ("Duplicate types: " ^ commas_quote dups);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    54
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    55
    val all_cons = map (Binding.name_of o first) (flat cons'');
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    56
    val test_dupl_cons =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    57
      case duplicates (op =) all_cons of 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    58
        [] => false | dups => error ("Duplicate constructors: " 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    59
                                      ^ commas_quote dups);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    60
    val all_sels =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    61
      (map Binding.name_of o map_filter second o maps second) (flat cons'');
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    62
    val test_dupl_sels =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    63
      case duplicates (op =) all_sels of
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    64
        [] => false | dups => error("Duplicate selectors: "^commas_quote dups);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    65
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    66
    fun test_dupl_tvars s =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    67
      case duplicates (op =) (map(fst o dest_TFree)s) of
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    68
        [] => false | dups => error("Duplicate type arguments: " 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    69
                                    ^commas_quote dups);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    70
    val test_dupl_tvars' = exists test_dupl_tvars (map snd dtnvs);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    71
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    72
    (* test for free type variables, illegal sort constraints on rhs,
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    73
       non-pcpo-types and invalid use of recursive type;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    74
       replace sorts in type variables on rhs *)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    75
    fun analyse_equation ((dname,typevars),cons') = 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    76
      let
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    77
        val tvars = map dest_TFree typevars;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    78
        val distinct_typevars = map TFree tvars;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    79
        fun rm_sorts (TFree(s,_)) = TFree(s,[])
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    80
          | rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    81
          | rm_sorts (TVar(s,_))  = TVar(s,[])
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    82
        and remove_sorts l = map rm_sorts l;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    83
        fun analyse indirect (TFree(v,s))  =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    84
            (case AList.lookup (op =) tvars v of 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    85
               NONE => error ("Free type variable " ^ quote v ^ " on rhs.")
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    86
             | SOME sort => if eq_set (op =) (s, defaultS) orelse
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    87
                               eq_set (op =) (s, sort)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    88
                            then TFree(v,sort)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    89
                            else error ("Inconsistent sort constraint" ^
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    90
                                        " for type variable " ^ quote v))
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    91
          | analyse indirect (t as Type(s,typl)) =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    92
            (case AList.lookup (op =) dtnvs s of
36119
ff605b8fdfdb remove dead code
huffman
parents: 36118
diff changeset
    93
               NONE => Type (s, map (analyse false) typl)
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    94
             | SOME typevars =>
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    95
                 if indirect 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    96
                 then error ("Indirect recursion of type " ^ 
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
    97
                             quote (string_of_typ thy t))
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    98
                 else if dname <> s orelse
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
    99
                         (** BUG OR FEATURE?:
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   100
                             mutual recursion may use different arguments **)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   101
                         remove_sorts typevars = remove_sorts typl 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   102
                 then Type(s,map (analyse true) typl)
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   103
                 else error ("Direct recursion of type " ^ 
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   104
                             quote (string_of_typ thy t) ^ 
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   105
                             " with different arguments"))
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   106
          | analyse indirect (TVar _) = Imposs "extender:analyse";
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   107
        fun check_pcpo lazy T =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   108
            let val sort = arg_sort lazy in
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   109
              if Sign.of_sort thy (T, sort) then T
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   110
              else error ("Constructor argument type is not of sort " ^
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   111
                          Syntax.string_of_sort_global thy sort ^ ": " ^
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   112
                          string_of_typ thy T)
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   113
            end;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   114
        fun analyse_arg (lazy, sel, T) =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   115
            (lazy, sel, check_pcpo lazy (analyse false T));
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   116
        fun analyse_con (b, args, mx) = (b, map analyse_arg args, mx);
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   117
      in ((dname,distinct_typevars), map analyse_con cons') end; 
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   118
  in ListPair.map analyse_equation (dtnvs,cons'')
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   119
  end; (* let *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
(* ----- calls for building new thy and thms -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   122
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   123
type info =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   124
     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
   125
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30916
diff changeset
   126
fun gen_add_domain
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   127
    (prep_typ : theory -> 'a -> typ)
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   128
    (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
   129
    (arg_sort : bool -> sort)
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   130
    (comp_dbind : binding)
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   131
    (eqs''' : ((string * string option) list * binding * mixfix *
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   132
               (binding * (bool * binding option * 'a) list * mixfix) list) list)
35516
59550ec4921d get rid of primes on thy variables
huffman
parents: 35498
diff changeset
   133
    (thy : theory) =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   134
  let
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   135
    val dtnvs : (binding * typ list * mixfix) list =
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   136
      let
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   137
        fun readS (SOME s) = Syntax.read_sort_global thy s
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   138
          | readS NONE = Sign.defaultS thy;
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   139
        fun readTFree (a, s) = TFree (a, readS s);
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   140
      in
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   141
        map (fn (vs,dname:binding,mx,_) =>
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   142
                (dname, map readTFree vs, mx)) eqs'''
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   143
      end;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   144
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   145
    fun thy_type  (dname,tvars,mx) = (dname, length tvars, mx);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   146
    fun thy_arity (dname,tvars,mx) =
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39965
diff changeset
   147
      (Sign.full_name thy dname, map (snd o dest_TFree) tvars, arg_sort false);
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   148
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   149
    (* this theory is used just for parsing and error checking *)
35516
59550ec4921d get rid of primes on thy variables
huffman
parents: 35498
diff changeset
   150
    val tmp_thy = thy
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   151
      |> Theory.copy
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   152
      |> Sign.add_types (map thy_type dtnvs)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   153
      |> fold (AxClass.axiomatize_arity o thy_arity) dtnvs;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   154
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   155
    val dbinds : binding list =
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   156
        map (fn (_,dbind,_,_) => dbind) eqs''';
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   157
    val cons''' :
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   158
        (binding * (bool * binding option * 'a) list * mixfix) list list =
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   159
        map (fn (_,_,_,cons) => cons) eqs''';
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   160
    val cons'' :
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   161
        (binding * (bool * binding option * typ) list * mixfix) list list =
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   162
        map (map (upd_second (map (upd_third (prep_typ tmp_thy))))) cons''';
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   163
    val dtnvs' : (string * typ list) list =
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   164
        map (fn (dname,vs,mx) => (Sign.full_name thy dname,vs)) dtnvs;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   165
    val eqs' : ((string * typ list) *
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   166
        (binding * (bool * binding option * typ) list * mixfix) list) list =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   167
        check_and_sort_domain arg_sort dtnvs' cons'' tmp_thy;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   168
    val dts : typ list = map (Type o fst) eqs';
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   169
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   170
    fun mk_arg_typ (lazy, dest_opt, T) = if lazy then mk_uT T else T;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   171
    fun mk_con_typ (bind, args, mx) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   172
        if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   173
    fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons);
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   174
    val repTs : typ list = map mk_eq_typ eqs';
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   175
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   176
    val iso_spec : (binding * mixfix * (typ * typ)) list =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   177
        map (fn ((dbind, _, mx), eq) => (dbind, mx, eq))
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   178
          (dtnvs ~~ (dts ~~ repTs));
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   179
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   180
    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
   181
35520
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   182
    val new_dts : (string * string list) list =
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   183
        map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs';
f433f18d4c41 cleaned up, added type annotations
huffman
parents: 35519
diff changeset
   184
    fun one_con (con,args,mx) : cons =
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   185
        (Binding.name_of con,  (* FIXME preverse binding (!?) *)
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   186
         ListPair.map (fn ((lazy,sel,tp),vn) =>
35519
abf45a91d24d remove unused selector field from type arg
huffman
parents: 35518
diff changeset
   187
           mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), vn))
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   188
                      (args, Datatype_Prop.make_tnames (map third args)));
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   189
    val eqs : eq list =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   190
        map (fn (dtnvs,cons') => (dtnvs, map one_con cons')) eqs';
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   191
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   192
    val ((rewss, take_rews), theorems_thy) =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   193
        thy
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   194
          |> fold_map (fn (((dbind, eq), (_,cs)), info) =>
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   195
                Domain_Theorems.theorems (eq, eqs) dbind cs info take_info)
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   196
             (dbinds ~~ eqs ~~ eqs' ~~ iso_infos)
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   197
          ||>> Domain_Theorems.comp_theorems (comp_dbind, eqs) take_info;
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   198
  in
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   199
    theorems_thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 36960
diff changeset
   200
      |> Global_Theory.add_thmss
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   201
           [((Binding.qualified true "rews" comp_dbind,
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   202
              flat rewss @ take_rews), [])]
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   203
      |> snd
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   204
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   205
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   206
fun define_isos (spec : (binding * mixfix * (typ * typ)) list) =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   207
  let
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   208
    fun prep (dbind, mx, (lhsT, rhsT)) =
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   209
      let val (dname, vs) = dest_Type lhsT;
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   210
      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
   211
  in
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   212
    Domain_Isomorphism.domain_isomorphism (map prep spec)
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   213
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   214
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   215
fun pcpo_arg lazy = if lazy then @{sort cpo} else @{sort pcpo};
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39965
diff changeset
   216
fun rep_arg lazy = @{sort sfp};
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   217
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   218
val add_domain =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   219
    gen_add_domain Sign.certify_typ Domain_Axioms.add_axioms pcpo_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   220
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   221
val add_new_domain =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   222
    gen_add_domain Sign.certify_typ define_isos rep_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   223
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   224
val add_domain_cmd =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   225
    gen_add_domain Syntax.read_typ_global Domain_Axioms.add_axioms pcpo_arg;
36118
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   226
413d6d1f0f6e share more code between definitional and axiomatic domain packages
huffman
parents: 36117
diff changeset
   227
val add_new_domain_cmd =
39965
da88519e6a0b new_domain emits proper error message when a constructor argument type does not have sort 'rep'
huffman
parents: 39557
diff changeset
   228
    gen_add_domain Syntax.read_typ_global define_isos rep_arg;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   229
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   230
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   231
(** outer syntax **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   232
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   233
val _ = Keyword.keyword "lazy";
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   234
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30916
diff changeset
   235
val dest_decl : (bool * binding option * string) parser =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   236
  Parse.$$$ "(" |-- Scan.optional (Parse.$$$ "lazy" >> K true) false --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   237
    (Parse.binding >> SOME) -- (Parse.$$$ "::" |-- Parse.typ)  --| Parse.$$$ ")" >> Parse.triple1
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   238
    || Parse.$$$ "(" |-- Parse.$$$ "lazy" |-- Parse.typ --| Parse.$$$ ")"
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   239
    >> (fn t => (true,NONE,t))
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   240
    || Parse.typ >> (fn t => (false,NONE,t));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   241
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   242
val cons_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   243
  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
   244
a3d2128cac92 allow infix declarations for type constructors defined with domain package
huffman
parents: 30915
diff changeset
   245
val domain_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   246
  (Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix) --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   247
    (Parse.$$$ "=" |-- Parse.enum1 "|" cons_decl);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   248
30916
a3d2128cac92 allow infix declarations for type constructors defined with domain package
huffman
parents: 30915
diff changeset
   249
val domains_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   250
  Scan.option (Parse.$$$ "(" |-- Parse.binding --| Parse.$$$ ")") --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   251
    Parse.and_list1 domain_decl;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   252
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   253
fun mk_domain
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   254
    (definitional : bool)
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   255
    (opt_name : binding option,
33796
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   256
     doms : ((((string * string option) list * binding) * mixfix) *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   257
             ((binding * (bool * binding option * string) list) * mixfix) list) list ) =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   258
  let
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   259
    val names = map (fn (((_, t), _), _) => Binding.name_of t) doms;
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   260
    val specs : ((string * string option) list * binding * mixfix *
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   261
                 (binding * (bool * binding option * string) list * mixfix) list) list =
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   262
        map (fn (((vs, t), mx), cons) =>
6442aa3773a2 clean up indentation
huffman
parents: 33038
diff changeset
   263
                (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms;
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   264
    val comp_dbind =
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   265
        case opt_name of NONE => Binding.name (space_implode "_" names)
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   266
                       | SOME s => s;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   267
  in
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   268
    if definitional 
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   269
    then add_new_domain_cmd comp_dbind specs
218e9766a848 replace some string arguments with bindings
huffman
parents: 35659
diff changeset
   270
    else add_domain_cmd comp_dbind specs
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   271
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   272
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   273
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   274
  Outer_Syntax.command "domain" "define recursive domains (HOLCF)"
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   275
    Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain false));
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   276
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33796
diff changeset
   277
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   278
  Outer_Syntax.command "new_domain" "define recursive domains (HOLCF)"
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36119
diff changeset
   279
    Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain true));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   280
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   281
end;