src/HOLCF/domain/interface.ML
author berghofe
Mon, 24 Jan 2005 17:59:48 +0100
changeset 15457 1fbd4aba46e3
parent 14981 e73f8140af78
child 15531 08c8dad8e399
permissions -rw-r--r--
Adapted to modified interface of PureThy.get_thm(s).
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: 2267
diff changeset
     1
(*  Title:      HOLCF/domain/interface.ML
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2267
diff changeset
     2
    ID:         $Id$
12030
wenzelm
parents: 6642
diff changeset
     3
    Author:     David von Oheimb
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2267
diff changeset
     4
12030
wenzelm
parents: 6642
diff changeset
     5
Parser for domain section.
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     6
*)
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
     7
(** BEWARE: this is still needed for old-style theories **)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     8
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     9
local 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    10
  open ThyParse;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    11
  open Domain_Library;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    12
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    13
(* --- generation of bindings for axioms and theorems in .thy.ML - *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    14
15457
1fbd4aba46e3 Adapted to modified interface of PureThy.get_thm(s).
berghofe
parents: 14981
diff changeset
    15
  fun get      dname name   = "get_thm thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", None)";
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    16
  fun gen_vals dname name   = "val "^ name ^ " = get_thm" ^ 
4709
d24168720303 Symbol.explode;
wenzelm
parents: 4122
diff changeset
    17
			(if hd (rev (Symbol.explode name)) = "s" then "s" else "")^
15457
1fbd4aba46e3 Adapted to modified interface of PureThy.get_thm(s).
berghofe
parents: 14981
diff changeset
    18
			" thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", None);";
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    19
  fun gen_vall       name l = "val "^name^" = "^ mk_list l^";";
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    20
  val rews = "iso_rews @ when_rews @\n\
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    21
 	   \  con_rews @ sel_rews @ dis_rews @ dist_les @ dist_eqs @\n\
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    22
	   \  copy_rews";
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    23
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    24
  fun gen_struct num ((dname,_), cons') = let
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    25
    val axioms1 = ["abs_iso", "rep_iso", "when_def"];
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    26
		   (* con_defs , sel_defs, dis_defs *) 
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    27
    val axioms2 = ["copy_def", "reach", "take_def", "finite_def"];
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    28
    val theorems = ["iso_rews", "exhaust", "casedist", "when_rews", "con_rews", 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    29
		    "sel_rews", "dis_rews", "dist_les", "dist_eqs", "inverts", 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    30
		    "injects", "copy_rews"];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    31
    in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    32
      "structure "^dname^" = struct\n"^
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    33
      cat_lines (map (gen_vals dname) axioms1)^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    34
      gen_vall "con_defs"       (map (fn (con,_,_) => 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    35
		get dname (strip_esc con^"_def")) cons')^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    36
      gen_vall "sel_defs" (flat (map (fn (_,_,args) => map (fn (_,sel,_) => 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    37
		get dname (sel^"_def")) args)    cons'))^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    38
      gen_vall "dis_defs"       (map (fn (con,_,_) => 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    39
		get dname (dis_name_ con^"_def")) cons')^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    40
      cat_lines (map (gen_vals dname) axioms2)^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    41
      cat_lines (map (gen_vals dname) theorems)^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    42
      (if num > 1 then "val rews = " ^rews ^";\n" else "")
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    43
    end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    44
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
    45
  fun mk_domain eqs'' = let
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
    46
    val num    = length eqs'';
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
    47
    val dtnvs  = map fst eqs'';
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    48
    val dnames = map fst dtnvs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    49
    val comp_dname = implode (separate "_" dnames);
2267
b2326aefecbc Replaced map...~~ by ListPair.map
paulson
parents: 2240
diff changeset
    50
    val conss' = ListPair.map 
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    51
	(fn (dnam,cons'') => let fun sel n m = upd_second 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    52
			     (fn None   => dnam^"_sel_"^(string_of_int n)^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    53
						    "_"^(string_of_int m)
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    54
			       | Some s => s)
2267
b2326aefecbc Replaced map...~~ by ListPair.map
paulson
parents: 2240
diff changeset
    55
	      fun fill_sels n con = upd_third (mapn (sel n) 1) con;
b2326aefecbc Replaced map...~~ by ListPair.map
paulson
parents: 2240
diff changeset
    56
	  in mapn fill_sels 1 cons'' end)
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
    57
	(dnames, map snd eqs'');
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    58
    val eqs' = dtnvs~~conss';
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    59
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    60
(* ----- string for calling add_domain -------------------------------------- *)
1284
e5b95ee2616b removed incompatibility with sml
regensbu
parents: 1274
diff changeset
    61
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    62
    val thy_ext_string = let
2240
a8c074224e11 Replaced obsolete "makestring" by Bool.toString
paulson
parents: 1637
diff changeset
    63
      fun mk_conslist cons' = 
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    64
	  mk_list (map (fn (c,syn,ts)=> "\n"^
14601
e46e7cdcf796 Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents: 14097
diff changeset
    65
		    mk_triple(Library.quote c, syn, mk_list (map (fn (b,s ,tp) =>
e46e7cdcf796 Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents: 14097
diff changeset
    66
		    mk_triple(Bool.toString b, Library.quote s, tp)) ts))) cons');
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    67
    in "val thy = thy |> Domain_Extender.add_domain "
14601
e46e7cdcf796 Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents: 14097
diff changeset
    68
       ^ mk_pair(Library.quote comp_dname, mk_list(map (fn ((n,vs),cs) => "\n"^
e46e7cdcf796 Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents: 14097
diff changeset
    69
				   mk_pair (mk_pair (Library.quote n, mk_list vs), 
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
    70
					    mk_conslist cs)) eqs'))
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    71
       ^ ";\n"
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    72
    end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    73
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    74
(* ----- string for producing the structures -------------------------------- *)
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    75
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    76
    val val_gen_string =  let
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    77
      val comp_axioms   = [(* copy, *) "reach", "take_def", "finite_def"
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    78
			   (*, "bisim_def" *)];
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    79
      val comp_theorems = ["take_rews", "take_lemmas", "finites", 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    80
			   "finite_ind", "ind", "coind"];
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    81
      fun collect sep name = if num = 1 then name else
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    82
		   implode (separate sep (map (fn s => s^"."^name) dnames));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    83
    in
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    84
	implode (separate "end; (* struct *)\n" 
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    85
			  (map (gen_struct num) eqs')) ^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    86
	(if num > 1 then "end; (* struct *)\n\
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    87
		         \structure "^comp_dname^" = struct\n" ^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    88
			 gen_vals comp_dname "copy_def" ^"\n" ^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    89
			 cat_lines (map (fn name => gen_vall (name^"s")
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    90
			 (map (fn dn => dn^"."^name) dnames)) comp_axioms)^"\n"
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    91
		    else "") ^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    92
	gen_vals comp_dname "bisim_def" ^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    93
        cat_lines (map (gen_vals comp_dname) comp_theorems)^"\n"^
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    94
	"val rews = "^(if num>1 then collect" @ " "rews"else rews)^ 
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    95
		    " @ take_rews;\n\
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    96
      \end; (* struct *)\n"
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    97
    end;
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
    98
  in (thy_ext_string^
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
    99
      val_gen_string^
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
   100
      "val thy = thy",
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
   101
      "") end;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   102
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
   103
(* ----- parser for domain declaration equation ----------------------------- *)
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   104
(** BEWARE: should be consistent with domains_decl in extender.ML **)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   105
6642
732af87c0650 strip_quotes replaced by unenclose;
wenzelm
parents: 4709
diff changeset
   106
  val name' = name >> unenclose;
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   107
  fun esc_quote s = let fun esc [] = []
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   108
			|   esc ("\""::xs) = esc xs
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   109
			|   esc ("[" ::xs) = "{"::esc xs
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   110
			|   esc ("]" ::xs) = "}"::esc xs
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   111
			|   esc (x   ::xs) = x  ::esc xs
4709
d24168720303 Symbol.explode;
wenzelm
parents: 4122
diff changeset
   112
		    in implode (esc (Symbol.explode s)) end;
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   113
  val type_var' = ((type_var >> unenclose) ^^ 
14601
e46e7cdcf796 Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents: 14097
diff changeset
   114
	          optional ($$ "::" ^^ !! (sort >> esc_quote)) "") >> Library.quote;
14097
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   115
  val type_args = (type_var' >> single
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   116
                 ||  "(" $$-- !! (list1 type_var' --$$ ")")
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   117
                 ||  empty >> K [])
f4d2ff3cad09 re-introduced sort constraints on LHS
oheimb
parents: 12030
diff changeset
   118
  val con_typ   = (type_args -- name' >> Library.swap)
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   119
                      (* here ident may be used instead of name' 
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   120
                         to avoid ambiguity with standard mixfix option *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   121
  val domain_arg  = "(" $$-- (optional ($$ "lazy" >> K true) false)
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2267
diff changeset
   122
			  -- (optional ((name' >> Some) --$$ "::") None)
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   123
			  -- typ --$$ ")" 
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   124
		    >> (fn ((lazy,sel),tp) => (lazy,sel,tp))
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3622
diff changeset
   125
		 || typ >> (fn x => (false,None,x)) 
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2267
diff changeset
   126
  val domain_cons = name' -- !! (repeat domain_arg) 
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   127
		    -- ThyParse.opt_mixfix
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   128
		    >> (fn ((con,args),syn) => (con,syn,args));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   129
in
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   130
  val domain_decl = (enum1 "and" (con_typ --$$ "="  -- !! 
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   131
				 (enum1 "|" domain_cons))) >> mk_domain;
3622
85898be702b2 use ThySyn.add_syntax;
wenzelm
parents: 2446
diff changeset
   132
  val _ = ThySyn.add_syntax
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   133
    ["lazy", "and"]
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
   134
    [("domain", domain_decl)]
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   135
3622
85898be702b2 use ThySyn.add_syntax;
wenzelm
parents: 2446
diff changeset
   136
end; (* local *)