author | wenzelm |
Thu, 23 Oct 1997 12:10:32 +0200 | |
changeset 3973 | 1be726ef6813 |
parent 3405 | 2cccd0e3e9ea |
child 4107 | 2270829d2364 |
permissions | -rw-r--r-- |
3302 | 1 |
(* Title: TFL/thry |
2 |
ID: $Id$ |
|
3 |
Author: Konrad Slind, Cambridge University Computer Laboratory |
|
4 |
Copyright 1997 University of Cambridge |
|
5 |
*) |
|
6 |
||
2112 | 7 |
structure Thry : Thry_sig (* LThry_sig *) = |
8 |
struct |
|
9 |
||
10 |
structure S = USyntax; |
|
11 |
||
12 |
fun THRY_ERR{func,mesg} = Utils.ERR{module = "Thry",func=func,mesg=mesg}; |
|
13 |
||
14 |
(*--------------------------------------------------------------------------- |
|
15 |
* Matching |
|
16 |
*---------------------------------------------------------------------------*) |
|
17 |
||
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3332
diff
changeset
|
18 |
local fun tybind (x,y) = (TVar (x,["term"]) , y) |
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3332
diff
changeset
|
19 |
fun tmbind (x,y) = (Var (x,type_of y), y) |
2112 | 20 |
in |
21 |
fun match_term thry pat ob = |
|
22 |
let val tsig = #tsig(Sign.rep_sg(sign_of thry)) |
|
23 |
val (ty_theta,tm_theta) = Pattern.match tsig (pat,ob) |
|
24 |
in (map tmbind tm_theta, map tybind ty_theta) |
|
25 |
end |
|
26 |
||
27 |
fun match_type thry pat ob = |
|
28 |
map tybind(Type.typ_match (#tsig(Sign.rep_sg(sign_of thry))) ([],(pat,ob))) |
|
29 |
end; |
|
30 |
||
31 |
||
32 |
(*--------------------------------------------------------------------------- |
|
33 |
* Typing |
|
34 |
*---------------------------------------------------------------------------*) |
|
35 |
||
36 |
fun typecheck thry = cterm_of (sign_of thry); |
|
37 |
||
38 |
||
39 |
(*--------------------------------------------------------------------------- |
|
40 |
* A collection of facts about datatypes |
|
41 |
*---------------------------------------------------------------------------*) |
|
42 |
val nat_record = Dtype.build_record (Nat.thy, ("nat",["0","Suc"]), nat_ind_tac) |
|
43 |
val prod_record = |
|
44 |
let val prod_case_thms = Dtype.case_thms (sign_of Prod.thy) [split] |
|
45 |
(fn s => res_inst_tac [("p",s)] PairE_lemma) |
|
46 |
fun const s = Const(s, the(Sign.const_type (sign_of Prod.thy) s)) |
|
47 |
in ("*", |
|
48 |
{constructors = [const "Pair"], |
|
49 |
case_const = const "split", |
|
50 |
case_rewrites = [split RS eq_reflection], |
|
51 |
case_cong = #case_cong prod_case_thms, |
|
52 |
nchotomy = #nchotomy prod_case_thms}) |
|
53 |
end; |
|
54 |
||
55 |
(*--------------------------------------------------------------------------- |
|
56 |
* Hacks to make interactive mode work. Referring to "datatypes" directly |
|
57 |
* is temporary, I hope! |
|
58 |
*---------------------------------------------------------------------------*) |
|
59 |
val match_info = fn thy => |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
60 |
fn "*" => Some({case_const = #case_const (#2 prod_record), |
2112 | 61 |
constructors = #constructors (#2 prod_record)}) |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
62 |
| "nat" => Some({case_const = #case_const (#2 nat_record), |
2112 | 63 |
constructors = #constructors (#2 nat_record)}) |
64 |
| ty => case assoc(!datatypes,ty) |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
65 |
of None => None |
2112 | 66 |
| Some{case_const,constructors, ...} => |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
67 |
Some{case_const=case_const, constructors=constructors} |
2112 | 68 |
|
69 |
val induct_info = fn thy => |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
70 |
fn "*" => Some({nchotomy = #nchotomy (#2 prod_record), |
2112 | 71 |
constructors = #constructors (#2 prod_record)}) |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
72 |
| "nat" => Some({nchotomy = #nchotomy (#2 nat_record), |
2112 | 73 |
constructors = #constructors (#2 nat_record)}) |
74 |
| ty => case assoc(!datatypes,ty) |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
75 |
of None => None |
2112 | 76 |
| Some{nchotomy,constructors, ...} => |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3191
diff
changeset
|
77 |
Some{nchotomy=nchotomy, constructors=constructors} |
2112 | 78 |
|
79 |
val extract_info = fn thy => |
|
80 |
let val case_congs = map (#case_cong o #2) (!datatypes) |
|
81 |
val case_rewrites = flat(map (#case_rewrites o #2) (!datatypes)) |
|
82 |
in {case_congs = #case_cong (#2 prod_record):: |
|
83 |
#case_cong (#2 nat_record)::case_congs, |
|
84 |
case_rewrites = #case_rewrites(#2 prod_record)@ |
|
85 |
#case_rewrites(#2 nat_record)@case_rewrites} |
|
86 |
end; |
|
87 |
||
88 |
end; (* Thry *) |