author | nipkow |
Fri, 08 Jul 1994 17:22:58 +0200 | |
changeset 92 | bcd0ee8d71aa |
parent 91 | a94029edb01f |
child 96 | d94d0b324b4b |
permissions | -rw-r--r-- |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
1 |
(* Title: HOL/Datatype |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
2 |
ID: $Id$ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
3 |
Author: Max Breitling / Carsten Clasohm |
53 | 4 |
Copyright 1994 TU Muenchen |
5 |
*) |
|
6 |
||
7 |
||
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
8 |
(*choice between Ci_neg1 and Ci_neg2 axioms depends on number of constructors*) |
92 | 9 |
local val dtK = 5 |
10 |
in |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
11 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
12 |
local open ThyParse in |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
13 |
val datatype_decls = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
14 |
let fun cat s1 s2 = s1 ^ " " ^ s2; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
15 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
16 |
val pars = parents "(" ")"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
17 |
val brackets = parents "[" "]"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
18 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
19 |
val mk_list = brackets o commas; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
20 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
21 |
val tvar = type_var >> cat "dtVar"; |
53 | 22 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
23 |
val type_var_list = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
24 |
tvar >> (fn s => [s]) || "(" $$-- list1 tvar --$$ ")"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
25 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
26 |
val typ = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
27 |
ident >> (cat "dtId" o quote) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
28 |
|| |
92 | 29 |
type_var_list -- ident >> (fn (ts, id) => "dtComp (" ^ mk_list ts ^ |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
30 |
", " ^ quote id ^ ")") |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
31 |
|| |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
32 |
tvar; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
33 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
34 |
val typ_list = "(" $$-- list1 typ --$$ ")" || empty; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
35 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
36 |
val cons = name -- typ_list -- opt_mixfix; |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
37 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
38 |
fun constructs ts = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
39 |
( cons --$$ "|" -- constructs >> op:: |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
40 |
|| |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
41 |
cons >> (fn c => [c])) ts; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
42 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
43 |
val mk_cons = map (fn ((s, ts), syn) => |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
44 |
pars (commas [s, mk_list ts, syn])); |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
45 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
46 |
(*remove all quotes from a string*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
47 |
fun rem_quotes s = implode (filter (fn c => c <> "\"") (explode s)); |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
48 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
49 |
(*generate names of ineq axioms*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
50 |
fun rules_ineq cs tname = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
51 |
let (*combine all constructor names with all others w/o duplicates*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
52 |
fun negOne _ [] = [] |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
53 |
| negOne (c : (string * 'a) * 'b) ((c2 : (string * 'a) * 'b) |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
54 |
:: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
55 |
quote ("ineq_" ^ rem_quotes (#1 (#1 c)) ^ "_" ^ |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
56 |
rem_quotes (#1 (#1 c2))) :: negOne c cs; |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
57 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
58 |
fun neg1 [] = [] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
59 |
| neg1 (c1 :: cs) = (negOne c1 cs) @ (neg1 cs) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
60 |
in if length cs < dtK then neg1 cs |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
61 |
else map (fn n => quote (tname ^ "_ord" ^ string_of_int n)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
62 |
(0 upto (length cs)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
63 |
end; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
64 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
65 |
fun arg1 ((_, ts), _) = not (null ts); |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
66 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
67 |
(*generate string for calling 'add_datatype'*) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
68 |
fun mk_params ((ts, tname), cons) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
69 |
("|> add_datatype\n" ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
70 |
pars (commas [mk_list ts, quote tname, mk_list (mk_cons cons)]), |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
71 |
"structure " ^ tname ^ " =\n\ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
72 |
\struct\n\ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
73 |
\ val inject = map (get_axiom thy) " ^ |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
74 |
mk_list (map (fn ((s,_), _) => quote ("inject_" ^ rem_quotes s)) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
75 |
(filter arg1 cons)) ^ ";\n\ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
76 |
\ val ineq = " ^ (if length cons < dtK then "let val ineq' = " else "") |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
77 |
^ "map (get_axiom thy) " ^ mk_list (rules_ineq cons tname) ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
78 |
(if length cons < dtK then |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
79 |
" in ineq' @ (map (fn t => sym COMP (t RS contrapos)) ineq') end" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
80 |
else "") ^ ";\n\ |
87 | 81 |
\ val induct = get_axiom thy \"" ^ tname ^ "_induct\";\n\ |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
82 |
\ val cases = map (get_axiom thy) " ^ |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
83 |
mk_list (map (fn ((s,_),_) => |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
84 |
quote(tname ^ "_case_" ^ rem_quotes s)) cons) ^ ";\n\ |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
85 |
\ val simps = inject @ ineq @ cases;\n\ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
86 |
\ fun induct_tac a = res_inst_tac [(" ^ quote tname ^ ", a)] induct;\n\ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
87 |
\end;\n"); |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
88 |
in (type_var_list || empty) -- ident --$$ "=" -- constructs >> mk_params end |
53 | 89 |
end; |
90 |
||
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
91 |
(*used for constructor parameters*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
92 |
datatype dt_type = dtVar of string | |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
93 |
dtId of string | |
92 | 94 |
dtComp of dt_type list * string | |
95 |
dtRek of dt_type list * string; |
|
53 | 96 |
|
92 | 97 |
local open Syntax.Mixfix |
98 |
exception Impossible |
|
99 |
in |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
100 |
fun add_datatype (typevars, tname, cons_list') thy = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
101 |
let fun cat s1 s2 = s1 ^ " " ^ s2; |
53 | 102 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
103 |
val pars = parents "(" ")"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
104 |
val brackets = parents "[" "]"; |
53 | 105 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
106 |
val mk_list = brackets o commas; |
53 | 107 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
108 |
(*check if constructor names are unique*) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
109 |
fun check_cons (cs : (string * 'b * 'c) list) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
110 |
(case findrep (map #1 cs) of |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
111 |
[] => true |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
112 |
| c::_ => error("Constructor \"" ^ c ^ "\" occurs twice")); |
53 | 113 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
114 |
(*search for free type variables and convert recursive *) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
115 |
fun analyse_types (cons, typlist, syn) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
116 |
let fun analyse ((dtVar v) :: typlist) = |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
117 |
if ((dtVar v) mem typevars) then |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
118 |
(dtVar v) :: analyse typlist |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
119 |
else error ("Variable " ^ v ^ " is free.") |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
120 |
| analyse ((dtId s) :: typlist) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
121 |
if tname<>s then (dtId s) :: analyse typlist |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
122 |
else if null typevars then |
92 | 123 |
dtRek ([], tname) :: analyse typlist |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
124 |
else error (s ^ " used in different ways") |
92 | 125 |
| analyse (dtComp (typl,s) :: typlist) = |
126 |
if tname <> s then dtComp (analyse typl, s) |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
127 |
:: analyse typlist |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
128 |
else if typevars = typl then |
92 | 129 |
dtRek (typl, s) :: analyse typlist |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
130 |
else |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
131 |
error (s ^ " used in different ways") |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
132 |
| analyse [] = [] |
92 | 133 |
| analyse ((dtRek _) :: _) = raise Impossible; |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
134 |
in (cons, analyse typlist, syn) end; |
53 | 135 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
136 |
(*test if there are elements that are not recursive, i.e. if the type is |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
137 |
not empty*) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
138 |
fun one_not_rek (cs : ('a * dt_type list * 'c) list) = |
92 | 139 |
let val contains_no_rek = forall (fn dtRek _ => false | _ => true); |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
140 |
in exists (contains_no_rek o #2) cs orelse |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
141 |
error("Empty type not allowed!") end; |
53 | 142 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
143 |
val dummy = check_cons cons_list'; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
144 |
val cons_list = map analyse_types cons_list'; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
145 |
val dummy = one_not_rek cons_list; |
53 | 146 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
147 |
(*Pretty printers for type lists; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
148 |
pp_typlist1: parentheses, pp_typlist2: brackets*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
149 |
fun pp_typ (dtVar s) = s |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
150 |
| pp_typ (dtId s) = s |
92 | 151 |
| pp_typ (dtComp (typvars, id)) = (pp_typlist1 typvars) ^ id |
152 |
| pp_typ (dtRek (typvars, id)) = (pp_typlist1 typvars) ^ id |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
153 |
and |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
154 |
pp_typlist' ts = commas (map pp_typ ts) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
155 |
and |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
156 |
pp_typlist1 ts = if null ts then "" else pars (pp_typlist' ts); |
53 | 157 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
158 |
fun pp_typlist2 ts = if null ts then "" else brackets (pp_typlist' ts); |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
159 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
160 |
fun Args(var, delim, n, m) = if n = m then var ^ string_of_int(n) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
161 |
else var ^ string_of_int(n) ^ delim ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
162 |
Args(var, delim, n+1, m); |
53 | 163 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
164 |
(* Generate syntax translation for case rules *) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
165 |
fun calc_xrules c_nr y_nr ((id, typlist, syn) :: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
166 |
let val name = const_name id syn; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
167 |
val arity = length typlist; |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
168 |
val body = "z" ^ string_of_int(c_nr); |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
169 |
val args1 = if arity=0 then "" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
170 |
else pars (Args ("y", ",", y_nr, y_nr+arity-1)); |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
171 |
val args2 = if arity=0 then "" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
172 |
else "% " ^ Args ("y", " ", y_nr, y_nr+arity-1) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
173 |
^ ". "; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
174 |
val (rest1,rest2) = |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
175 |
if null cs then ("","") |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
176 |
else let val (h1, h2) = calc_xrules (c_nr+1) (y_nr+arity) cs |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
177 |
in (" | " ^ h1, ", " ^ h2) end; |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
178 |
in (name ^ args1 ^ " => " ^ body ^ rest1, args2 ^ body ^ rest2) end |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
179 |
| calc_xrules _ _ [] = raise Impossible; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
180 |
|
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
181 |
val xrules = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
182 |
let val (first_part, scnd_part) = calc_xrules 1 1 cons_list |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
183 |
in [("logic", "case x of " ^ first_part) <-> |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
184 |
("logic", tname ^ "_case(x, " ^ scnd_part ^ ")" )] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
185 |
end; |
53 | 186 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
187 |
(*type declarations for constructors*) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
188 |
fun const_types ((id, typlist, syn) :: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
189 |
(id, |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
190 |
(if null typlist then "" else pp_typlist2 typlist ^ " => ") ^ |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
191 |
pp_typlist1 typevars ^ tname, syn) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
192 |
:: const_types cs |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
193 |
| const_types [] = []; |
53 | 194 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
195 |
fun create_typevar (dtVar s) typlist = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
196 |
if (dtVar s) mem typlist then |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
197 |
create_typevar (dtVar (s ^ "'")) typlist |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
198 |
else s |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
199 |
| create_typevar _ _ = raise Impossible; |
53 | 200 |
|
92 | 201 |
fun assumpt (dtRek _ :: ts, v :: vs ,found) = |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
202 |
let val h = if found then ";P(" ^ v ^ ")" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
203 |
else "[| P(" ^ v ^ ")" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
204 |
in h ^ (assumpt (ts, vs, true)) end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
205 |
| assumpt (t :: ts, v :: vs, found) = assumpt (ts, vs, found) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
206 |
| assumpt ([], [], found) = if found then "|] ==>" else "" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
207 |
| assumpt _ = raise Impossible; |
53 | 208 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
209 |
(*insert type with suggested name 'varname' into table*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
210 |
fun insert typ varname ((t, s, n) :: xs) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
211 |
if typ = t then (t, s, n+1) :: xs |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
212 |
else if varname = s then (t,s,n) :: (insert typ (varname ^ "'") xs) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
213 |
else (t,s,n) :: (insert typ varname xs) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
214 |
| insert typ varname [] = [(typ, varname, 1)]; |
53 | 215 |
|
92 | 216 |
fun insert_types (dtRek (l,id) :: ts) tab = |
217 |
insert_types ts (insert (dtRek(l,id)) id tab) |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
218 |
| insert_types ((dtVar s) :: ts) tab = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
219 |
insert_types ts (insert (dtVar s) (implode (tl (explode s))) tab) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
220 |
| insert_types ((dtId s) :: ts) tab = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
221 |
insert_types ts (insert (dtId s) s tab) |
92 | 222 |
| insert_types (dtComp (l,id) :: ts) tab = |
223 |
insert_types ts (insert (dtComp(l,id)) id tab) |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
224 |
| insert_types [] tab = tab; |
53 | 225 |
|
92 | 226 |
fun update(dtRek _, s, v :: vs, (dtRek _) :: ts) = s :: vs |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
227 |
| update(t, s, v :: vs, t1 :: ts) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
228 |
if t=t1 then s :: vs |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
229 |
else v :: (update (t, s, vs, ts)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
230 |
| update _ = raise Impossible; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
231 |
|
92 | 232 |
fun update_n (dtRek r1, s, v :: vs, (dtRek r2) :: ts, n) = |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
233 |
if r1 = r2 then (s ^ string_of_int n) :: |
92 | 234 |
(update_n (dtRek r1, s, vs, ts, n+1)) |
235 |
else v :: (update_n (dtRek r1, s, vs, ts, n)) |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
236 |
| update_n (t, s, v :: vs, t1 :: ts, n) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
237 |
if t = t1 then (s ^ string_of_int n) :: |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
238 |
(update_n (t, s, vs, ts, n+1)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
239 |
else v :: (update_n (t, s, vs, ts, n)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
240 |
| update_n (_,_,[],[],_) = [] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
241 |
| update_n _ = raise Impossible; |
53 | 242 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
243 |
(*insert type variables into table*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
244 |
fun convert ((t, s, n) :: ts) var_list typ_list = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
245 |
let val h = if n=1 then update (t, s, var_list, typ_list) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
246 |
else update_n (t, s, var_list, typ_list, 1) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
247 |
in convert ts h typ_list end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
248 |
| convert [] var_list _ = var_list; |
53 | 249 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
250 |
fun empty_list n = replicate n ""; |
53 | 251 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
252 |
fun t_inducting ((id, typl, syn) :: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
253 |
let val name = const_name id syn; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
254 |
val tab = insert_types typl []; |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
255 |
val arity = length typl; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
256 |
val var_list = convert tab (empty_list arity) typl; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
257 |
val h = if arity = 0 then " P(" ^ name ^ ")" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
258 |
else " !!" ^ (space_implode " " var_list) ^ "." ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
259 |
(assumpt (typl, var_list, false)) ^ "P(" ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
260 |
name ^ "(" ^ (commas var_list) ^ "))"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
261 |
val rest = t_inducting cs; |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
262 |
in if rest = "" then h else h ^ "; " ^ rest end |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
263 |
| t_inducting [] = ""; |
53 | 264 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
265 |
fun t_induct cl typ_name= |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
266 |
"[|" ^ t_inducting cl ^ "|] ==> P(" ^ typ_name ^ ")"; |
53 | 267 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
268 |
fun case_typlist typevar ((_, typlist, _) :: cs) = |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
269 |
let val h = if (length typlist) > 0 then |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
270 |
(pp_typlist2 typlist) ^ "=>" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
271 |
else "" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
272 |
in "," ^ h ^ typevar ^ (case_typlist typevar cs) end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
273 |
| case_typlist _ [] = ""; |
53 | 274 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
275 |
fun case_rules t_case arity n ((id, typlist, syn) :: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
276 |
let val name = const_name id syn; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
277 |
val args = if null typlist then "" |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
278 |
else "(" ^ Args ("x", ",", 1, length typlist) ^ ")" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
279 |
in (t_case ^ "_" ^ id, |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
280 |
t_case ^ "(" ^ name ^ args ^ "," ^ Args ("f", ",", 1, arity) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
281 |
^ ") = f" ^ string_of_int(n) ^ args) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
282 |
:: (case_rules t_case arity (n+1) cs) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
283 |
end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
284 |
| case_rules _ _ _ [] = []; |
53 | 285 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
286 |
val datatype_arity = length typevars; |
53 | 287 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
288 |
val types = [(tname, datatype_arity, NoSyn)]; |
53 | 289 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
290 |
val arities = |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
291 |
let val term_list = replicate datatype_arity ["term"]; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
292 |
in [(tname, term_list, ["term"])] end; |
53 | 293 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
294 |
val datatype_name = pp_typlist1 typevars ^ tname; |
53 | 295 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
296 |
val (case_const, rules_case) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
297 |
let val typevar = create_typevar (dtVar "'beta") typevars; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
298 |
val t_case = tname ^ "_case"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
299 |
val arity = length cons_list; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
300 |
val dekl = (t_case, "[" ^ pp_typlist1 typevars ^ tname ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
301 |
case_typlist typevar cons_list ^ "]=>" ^ typevar, NoSyn) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
302 |
:: nil; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
303 |
val rules = case_rules t_case arity 1 cons_list; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
304 |
in (dekl, rules) end; |
53 | 305 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
306 |
val consts = |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
307 |
const_types cons_list |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
308 |
@ (if length cons_list < dtK then [] |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
309 |
else [(tname ^ "_ord", datatype_name ^ "=>nat", NoSyn)]) |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
310 |
@ case_const; |
53 | 311 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
312 |
(*generate 'var_n, ..., var_m'*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
313 |
fun Args(var, delim, n, m) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
314 |
space_implode delim (map (fn n => var^string_of_int(n)) (n upto m)); |
53 | 315 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
316 |
(*generate 'name_1', ..., 'name_n'*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
317 |
fun C_exp(name, n, var) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
318 |
if n > 0 then name ^ "(" ^ Args (var, ",", 1, n) ^ ")" |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
319 |
else name; |
53 | 320 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
321 |
(*generate 'x_n = y_n, ..., x_m = y_m'*) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
322 |
fun Arg_eql(n,m) = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
323 |
if n=m then "x" ^ string_of_int(n) ^ "=y" ^ string_of_int(n) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
324 |
else "x" ^ string_of_int(n) ^ "=y" ^ string_of_int(n) ^ " & " ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
325 |
Arg_eql(n+1, m); |
53 | 326 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
327 |
fun Ci_ing ((id, typlist, syn) :: cs) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
328 |
let val name = const_name id syn; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
329 |
val arity = length typlist; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
330 |
in if arity > 0 |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
331 |
then ("inject_" ^ id, |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
332 |
"(" ^ C_exp(name,arity,"x") ^ "=" ^ C_exp(name,arity,"y") |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
333 |
^ ") = (" ^ Arg_eql (1, arity) ^ ")") :: (Ci_ing cs) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
334 |
else (Ci_ing cs) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
335 |
end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
336 |
| Ci_ing [] = []; |
53 | 337 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
338 |
fun Ci_negOne _ [] = [] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
339 |
| Ci_negOne c (c1::cs) = |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
340 |
let val (id1, tl1, syn1) = c |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
341 |
val (id2, tl2, syn2) = c1 |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
342 |
val name1 = const_name id1 syn1; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
343 |
val name2 = const_name id2 syn2; |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
344 |
val arit1 = length tl1 |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
345 |
val arit2 = length tl2 |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
346 |
val h = "(" ^ C_exp(name1, arit1, "x") ^ "~=" ^ |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
347 |
C_exp(name2, arit2, "y") ^ ")" |
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
348 |
in ("ineq_" ^ id1 ^ "_" ^ id2, h):: (Ci_negOne c cs) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
349 |
end; |
53 | 350 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
351 |
fun Ci_neg1 [] = [] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
352 |
| Ci_neg1 (c1::cs) = Ci_negOne c1 cs @ Ci_neg1 cs; |
53 | 353 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
354 |
fun suc_expr n = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
355 |
if n=0 then "0" else "Suc(" ^ suc_expr(n-1) ^ ")"; |
53 | 356 |
|
91
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
357 |
fun Ci_neg2equals (ord_t, ((id, typlist, syn) :: cs), n) = |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
358 |
let val name = const_name id syn; |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
359 |
val h = ord_t ^ "(" ^ (C_exp(name, length typlist, "x")) |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
360 |
^ ") = " ^ (suc_expr n) |
a94029edb01f
added mixfix annotations to constructor declarations
clasohm
parents:
87
diff
changeset
|
361 |
in (ord_t ^ (string_of_int (n+1)), h) |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
362 |
:: (Ci_neg2equals (ord_t, cs , n+1)) |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
363 |
end |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
364 |
| Ci_neg2equals (_, [], _) = []; |
53 | 365 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
366 |
val Ci_neg2 = |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
367 |
let val ord_t = tname ^ "_ord"; |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
368 |
in (Ci_neg2equals (ord_t, cons_list, 0)) @ |
87 | 369 |
[(ord_t ^ "0", |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
370 |
"(" ^ ord_t ^ "(x) ~= " ^ ord_t ^ "(y)) ==> (x ~= y)")] |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
371 |
end; |
53 | 372 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
373 |
val rules_ineq = if length cons_list < dtK then Ci_neg1 cons_list |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
374 |
else Ci_neg2; |
53 | 375 |
|
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
376 |
val rules_inject = Ci_ing cons_list; |
53 | 377 |
|
87 | 378 |
val rule_induct = (tname ^ "_induct", t_induct cons_list tname); |
60 | 379 |
|
87 | 380 |
val rules = rule_induct :: (rules_inject @ rules_ineq @ rules_case); |
80
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
381 |
in thy |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
382 |
|> add_types types |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
383 |
|> add_arities arities |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
384 |
|> add_consts consts |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
385 |
|> add_trrules xrules |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
386 |
|> add_axioms rules |
d3d727449d7b
datatypes must now be defined using a thy file section
clasohm
parents:
60
diff
changeset
|
387 |
end |
92 | 388 |
end |
389 |
end; |