src/HOLCF/domain/extender.ML
author haftmann
Mon, 12 Sep 2005 18:20:32 +0200
changeset 17325 d9d50222808e
parent 17057 0934ac31985f
child 18083 cf7669049df5
permissions -rw-r--r--
introduced new-style AList operations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
     1
(*  Title:      HOLCF/domain/extender.ML
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
     2
    ID:         $Id$
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
     3
    Author:     David von Oheimb
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
     4
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
     5
Theory extender for domain section, including new-style theory syntax.
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
     6
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
     7
###TODO: 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
     8
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
     9
this definition
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    10
domain empty = silly empty
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    11
yields
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    12
Exception-
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    13
   TERM
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    14
      ("typ_of_term: bad encoding of type",
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15187
diff changeset
    15
         [Abs ("uu", "_", Const ("NONE", "_"))]) raised
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    16
but this works fine:
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    17
domain Empty = silly Empty
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    18
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    19
strange syntax errors are produced for:
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    20
domain xx = xx ("x yy")
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    21
domain 'a foo = foo (sel::"'a") 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    22
and bar = bar ("'a dummy")
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    23
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    24
*)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    25
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    26
signature DOMAIN_EXTENDER =
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    27
sig
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    28
  val add_domain: string *
16394
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
    29
      ((bstring * string list) * (string * mixfix * (bool * string option * string) list) list) list
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    30
    -> theory -> theory
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    31
end;
2446
c2a9bf6c0948 The previous log message was wrong. The correct one is:
oheimb
parents: 2445
diff changeset
    32
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
    33
structure Domain_Extender: DOMAIN_EXTENDER =
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    34
struct
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    35
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    36
open Domain_Library;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    37
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
    38
(* ----- general testing and preprocessing of constructor list -------------- *)
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    39
  fun check_and_sort_domain (dtnvs: (string * typ list) list, 
16394
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
    40
     cons'' : ((string * mixfix * (bool*string option*typ) list) list) list) sg =
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    41
  let
7653
0408848fa7d4 Sign.defaultS;
wenzelm
parents: 4753
diff changeset
    42
    val defaultS = Sign.defaultS sg;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    43
    val test_dupl_typs = (case duplicates (map fst dtnvs) of 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    44
	[] => false | dups => error ("Duplicate types: " ^ commas_quote dups));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    45
    val test_dupl_cons = (case duplicates (map first (List.concat cons'')) of 
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    46
	[] => false | dups => error ("Duplicate constructors: " 
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    47
							 ^ commas_quote dups));
16394
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
    48
    val test_dupl_sels = (case duplicates (List.mapPartial second
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
    49
			       (List.concat (map third (List.concat cons'')))) of
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
    50
        [] => false | dups => error("Duplicate selectors: "^commas_quote dups));
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
    51
    val test_dupl_tvars = exists(fn s=>case duplicates(map(fst o rep_TFree)s)of
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 1637
diff changeset
    52
	[] => false | dups => error("Duplicate type arguments: " 
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    53
		   ^commas_quote dups)) (map snd dtnvs);
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    54
    (* test for free type variables, illegal sort constraints on rhs,
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    55
	       non-pcpo-types and invalid use of recursive type;
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    56
       replace sorts in type variables on rhs *)
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    57
    fun analyse_equation ((dname,typevars),cons') = 
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    58
      let
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    59
	val tvars = map rep_TFree typevars;
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    60
	fun distinct_name s = "'"^Sign.base_name dname^"_"^s;
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    61
	val distinct_typevars = map (fn (n,sort) => 
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    62
				     TFree (distinct_name n,sort)) tvars;
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    63
	fun rm_sorts (TFree(s,_)) = TFree(s,[])
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    64
	|   rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts)
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    65
	|   rm_sorts (TVar(s,_))  = TVar(s,[])
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    66
	and remove_sorts l = map rm_sorts l;
15601
2de79f493856 domain package now permits indirect recursion with these type constructors: *, ->, ++, **, u
huffman
parents: 15570
diff changeset
    67
	val indirect_ok = ["*","Cfun.->","Ssum.++","Sprod.**","Up.u"]
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17057
diff changeset
    68
	fun analyse indirect (TFree(v,s))  = (case AList.lookup (op =) tvars v of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15187
diff changeset
    69
		    NONE => error ("Free type variable " ^ quote v ^ " on rhs.")
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15187
diff changeset
    70
	          | SOME sort => if eq_set_string (s,defaultS) orelse
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    71
				    eq_set_string (s,sort    )
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    72
				 then TFree(distinct_name v,sort)
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
    73
				 else error ("Inconsistent sort constraint" ^
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
    74
				             " for type variable " ^ quote v))
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17057
diff changeset
    75
        |   analyse indirect (t as Type(s,typl)) = (case AList.lookup (op =) dtnvs s of
15601
2de79f493856 domain package now permits indirect recursion with these type constructors: *, ->, ++, **, u
huffman
parents: 15570
diff changeset
    76
		NONE          => if exists (fn x => x = s) indirect_ok
2de79f493856 domain package now permits indirect recursion with these type constructors: *, ->, ++, **, u
huffman
parents: 15570
diff changeset
    77
				 then Type(s,map (analyse false) typl)
2de79f493856 domain package now permits indirect recursion with these type constructors: *, ->, ++, **, u
huffman
parents: 15570
diff changeset
    78
				 else Type(s,map (analyse true) typl)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15187
diff changeset
    79
	      | SOME typevars => if indirect 
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    80
                           then error ("Indirect recursion of type " ^ 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    81
				        quote (string_of_typ sg t))
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    82
                           else if dname <> s orelse (** BUG OR FEATURE?: 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    83
                                mutual recursion may use different arguments **)
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    84
				   remove_sorts typevars = remove_sorts typl 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    85
				then Type(s,map (analyse true) typl)
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    86
				else error ("Direct recursion of type " ^ 
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    87
					     quote (string_of_typ sg t) ^ 
4753
b3aab5c73b52 improved checks
oheimb
parents: 4709
diff changeset
    88
					    " with different arguments"))
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    89
        |   analyse indirect (TVar _) = Imposs "extender:analyse";
4753
b3aab5c73b52 improved checks
oheimb
parents: 4709
diff changeset
    90
	fun check_pcpo t = (pcpo_type sg t orelse error(
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
    91
      "Constructor argument type is not of sort pcpo: "^string_of_typ sg t); t);
15187
8b74a39dba89 added check against indirect recursion
oheimb
parents: 14981
diff changeset
    92
	val analyse_con = upd_third (map (upd_third (check_pcpo o analyse false)));
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    93
      in ((dname,distinct_typevars), map analyse_con cons') end; 
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    94
  in ListPair.map analyse_equation (dtnvs,cons'')
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    95
  end; (* let *)
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    96
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    97
(* ----- calls for building new thy and thms -------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    98
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
    99
  fun add_domain (comp_dnam,eqs''') thy''' = let
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   100
    val sg''' = sign_of thy''';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   101
    val dtnvs = map ((fn (dname,vs) => 
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   102
			 (Sign.full_name sg''' dname,map (str2typ sg''') vs))
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   103
                   o fst) eqs''';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   104
    val cons''' = map snd eqs''';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   105
    fun thy_type  (dname,tvars)  = (Sign.base_name dname, length tvars, NoSyn);
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   106
    fun thy_arity (dname,tvars)  = (dname, map (snd o rep_TFree) tvars, pcpoS);
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   107
    val thy'' = thy''' |> Theory.add_types     (map thy_type  dtnvs)
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   108
		       |> Theory.add_arities_i (map thy_arity dtnvs);
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   109
    val sg'' = sign_of thy'';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   110
    val cons''=map (map (upd_third (map (upd_third (str2typ sg''))))) cons''';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   111
    val eqs' = check_and_sort_domain (dtnvs,cons'') sg'';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   112
    val thy' = thy'' |> Domain_Syntax.add_syntax (comp_dnam,eqs');
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   113
    val dts  = map (Type o fst) eqs';
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   114
    fun strip ss = Library.drop (find_index_eq "'" ss +1, ss);
11728
b5f6963b193c fixed typid;
wenzelm
parents: 7653
diff changeset
   115
    fun typid (Type  (id,_)) =
b5f6963b193c fixed typid;
wenzelm
parents: 7653
diff changeset
   116
          let val c = hd (Symbol.explode (Sign.base_name id))
b5f6963b193c fixed typid;
wenzelm
parents: 7653
diff changeset
   117
          in if Symbol.is_letter c then c else "t" end
b5f6963b193c fixed typid;
wenzelm
parents: 7653
diff changeset
   118
      | typid (TFree (id,_)   ) = hd (strip (tl (Symbol.explode id)))
b5f6963b193c fixed typid;
wenzelm
parents: 7653
diff changeset
   119
      | typid (TVar ((id,_),_)) = hd (tl (Symbol.explode id));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   120
    fun cons cons' = (map (fn (con,syn,args) =>
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   121
	((Syntax.const_name con syn),
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 2446
diff changeset
   122
	 ListPair.map (fn ((lazy,sel,tp),vn) => ((lazy,
4188
1025a27b08f9 changed libraray function find to find_index_eq, currying it
oheimb
parents: 4122
diff changeset
   123
					find_index_eq tp dts),
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   124
					sel,vn))
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   125
	     (args,(mk_var_names(map (typid o third) args)))
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   126
	 )) cons') : cons list;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   127
    val eqs = map (fn (dtnvs,cons') => (dtnvs,cons cons')) eqs' : eq list;
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   128
    val thy        = thy' |> Domain_Axioms.add_axioms (comp_dnam,eqs);
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   129
    val (theorems_thy, (rewss, take_rews)) = (foldl_map (fn (thy0,eq) =>
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   130
      Domain_Theorems.theorems (eq,eqs) thy0) (thy,eqs))
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   131
      |>>> Domain_Theorems.comp_theorems (comp_dnam, eqs);
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   132
in
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   133
  theorems_thy
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   134
  |> Theory.add_path (Sign.base_name comp_dnam)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   135
  |> (#1 o (PureThy.add_thmss [(("rews", List.concat rewss @ take_rews), [])]))
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   136
  |> Theory.parent_path
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   137
end;
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   138
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   139
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   140
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   141
(** outer syntax **)
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   142
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 16394
diff changeset
   143
local structure P = OuterParse and K = OuterKeyword in
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   144
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   145
val dest_decl =
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   146
  P.$$$ "(" |-- Scan.optional (P.$$$ "lazy" >> K true) false --
16394
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
   147
    (P.name >> SOME) -- (P.$$$ "::" |-- P.typ)  --| P.$$$ ")" >> P.triple1
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
   148
  || P.$$$ "(" |-- P.$$$ "lazy" |-- P.typ --| P.$$$ ")"
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
   149
       >> (fn t => (true,NONE,t))
495dbcd4f4c9 in domain declarations, selector names are now optional
huffman
parents: 16070
diff changeset
   150
  || P.typ >> (fn t => (false,NONE,t));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   151
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   152
val cons_decl =
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12037
diff changeset
   153
  P.name -- Scan.repeat dest_decl -- P.opt_mixfix
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   154
  >> (fn ((c, ds), mx) => (c, mx, ds));
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   155
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   156
val type_var' = (P.type_ident ^^ 
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   157
                 Scan.optional (P.$$$ "::" ^^ P.!!! P.sort) "");
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   158
val type_args' = type_var' >> single ||
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   159
                 P.$$$ "(" |-- P.!!! (P.list1 type_var' --| P.$$$ ")") ||
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   160
 		 Scan.succeed [];
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   161
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   162
val domain_decl = (type_args' -- P.name >> Library.swap) -- 
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12876
diff changeset
   163
                  (P.$$$ "=" |-- P.enum1 "|" cons_decl);
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   164
val domains_decl =
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   165
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.and_list1 domain_decl
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   166
  >> (fn (opt_name, doms) =>
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15187
diff changeset
   167
      (case opt_name of NONE => space_implode "_" (map (#1 o #1) doms) | SOME s => s, doms));
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   168
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   169
val domainP =
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   170
  OuterSyntax.command "domain" "define recursive domains (HOLCF)" K.thy_decl
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   171
    (domains_decl >> (Toplevel.theory o add_domain));
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   172
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   173
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   174
val _ = OuterSyntax.add_keywords ["lazy"];
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   175
val _ = OuterSyntax.add_parsers [domainP];
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   176
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   177
end;
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   178
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   179
end;