TFL/thry.sml
author paulson
Tue, 27 May 1997 13:22:30 +0200
changeset 3353 9112a2efb9a3
parent 3332 3921ebbd9cf0
child 3388 dbf61e36f8e9
permissions -rw-r--r--
Removal of module Mask and datatype binding with its constructor |->
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     1
(*  Title:      TFL/thry
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     2
    ID:         $Id$
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     4
    Copyright   1997  University of Cambridge
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     5
*)
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     6
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     7
structure Thry : Thry_sig (* LThry_sig *) = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     8
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     9
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
structure USyntax  = USyntax;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    11
structure S = USyntax;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    12
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    13
fun THRY_ERR{func,mesg} = Utils.ERR{module = "Thry",func=func,mesg=mesg};
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    15
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    16
 *    Matching 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    17
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    18
3353
9112a2efb9a3 Removal of module Mask and datatype binding with its constructor |->
paulson
parents: 3332
diff changeset
    19
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
    20
      fun tmbind (x,y) = (Var  (x,type_of y), y)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    21
in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    22
 fun match_term thry pat ob = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    23
    let val tsig = #tsig(Sign.rep_sg(sign_of thry))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    24
        val (ty_theta,tm_theta) = Pattern.match tsig (pat,ob)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    25
    in (map tmbind tm_theta, map tybind ty_theta)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    26
    end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    27
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    28
 fun match_type thry pat ob = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
    map tybind(Type.typ_match (#tsig(Sign.rep_sg(sign_of thry))) ([],(pat,ob)))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    31
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    32
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    33
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    34
 * Typing 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    35
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    36
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    37
fun typecheck thry = cterm_of (sign_of thry);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    38
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    39
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    40
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    41
(*----------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    42
 * Making a definition. The argument "tm" looks like "f = WFREC R M". This 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    43
 * entrypoint is specialized for interactive use, since it closes the theory
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    44
 * after making the definition. This allows later interactive definitions to
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    45
 * refer to previous ones. The name for the new theory is automatically 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    46
 * generated from the name of the argument theory.
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    47
 *---------------------------------------------------------------------------*)
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    48
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    49
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    50
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    51
 * TFL attempts to make definitions where the lhs is a variable. Isabelle
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    52
 * wants it to be a constant, so here we map it to a constant. Moreover, the
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    53
 * theory should already have the constant, so we refrain from adding the
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    54
 * constant to the theory. We just add the axiom and return the theory.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    55
 *---------------------------------------------------------------------------*)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    56
local val (imp $ tprop $ (eeq $ _ $ _ )) = #prop(rep_thm(eq_reflection))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    57
      val Const(eeq_name, ty) = eeq
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    58
      val prop = #2 (S.strip_type ty)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    59
in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    60
fun make_definition parent s tm = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    61
   let val {lhs,rhs} = S.dest_eq tm
3332
3921ebbd9cf0 More de-HOL-ification
paulson
parents: 3302
diff changeset
    62
       val (Name,Ty) = dest_Free lhs
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    63
       val lhs1 = S.mk_const{Name = Name, Ty = Ty}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    64
       val eeq1 = S.mk_const{Name = eeq_name, Ty = Ty --> Ty --> prop}
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    65
       val dtm = list_comb(eeq1,[lhs1,rhs])      (* Rename "=" to "==" *)
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    66
       val (_, tm', _) = Sign.infer_types (sign_of parent)
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    67
                     (K None) (K None) [] true ([dtm],propT)
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    68
       val new_thy = add_defs_i [(s,tm')] parent
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    69
   in 
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    70
   (freezeT((get_axiom new_thy s) RS meta_eq_to_obj_eq), new_thy)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    71
   end;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    72
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    73
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    74
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    75
 * Utility routine. Insert into list ordered by the key (a string). If two 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    76
 * keys are equal, the new element replaces the old. A more efficient option 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    77
 * for the future is needed. In fact, having the list of datatype facts be 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    78
 * ordered is useless, since the lookup should never fail!
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    79
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    80
fun insert (el as (x:string, _)) = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    81
 let fun canfind[] = [el] 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    82
       | canfind(alist as ((y as (k,_))::rst)) = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    83
           if (x<k) then el::alist
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    84
           else if (x=k) then el::rst
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    85
           else y::canfind rst 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    86
 in canfind
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    87
 end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    88
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    89
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    90
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    91
 *     A collection of facts about datatypes
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    92
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    93
val nat_record = Dtype.build_record (Nat.thy, ("nat",["0","Suc"]), nat_ind_tac)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    94
val prod_record =
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    95
    let val prod_case_thms = Dtype.case_thms (sign_of Prod.thy) [split] 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    96
                                 (fn s => res_inst_tac [("p",s)] PairE_lemma)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    97
         fun const s = Const(s, the(Sign.const_type (sign_of Prod.thy) s))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    98
     in ("*", 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    99
         {constructors = [const "Pair"],
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   100
            case_const = const "split",
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   101
         case_rewrites = [split RS eq_reflection],
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   102
             case_cong = #case_cong prod_case_thms,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   103
              nchotomy = #nchotomy prod_case_thms}) 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   104
     end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   105
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   106
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   107
 * Hacks to make interactive mode work. Referring to "datatypes" directly
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   108
 * is temporary, I hope!
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   109
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   110
val match_info = fn thy =>
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   111
    fn "*" => Some({case_const = #case_const (#2 prod_record),
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   112
                     constructors = #constructors (#2 prod_record)})
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   113
     | "nat" => Some({case_const = #case_const (#2 nat_record),
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   114
                       constructors = #constructors (#2 nat_record)})
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   115
     | ty => case assoc(!datatypes,ty)
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   116
               of None => None
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   117
                | Some{case_const,constructors, ...} =>
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   118
                   Some{case_const=case_const, constructors=constructors}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   119
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   120
val induct_info = fn thy =>
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   121
    fn "*" => Some({nchotomy = #nchotomy (#2 prod_record),
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   122
                     constructors = #constructors (#2 prod_record)})
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   123
     | "nat" => Some({nchotomy = #nchotomy (#2 nat_record),
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   124
                       constructors = #constructors (#2 nat_record)})
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   125
     | ty => case assoc(!datatypes,ty)
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   126
               of None => None
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   127
                | Some{nchotomy,constructors, ...} =>
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
   128
                  Some{nchotomy=nchotomy, constructors=constructors}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   129
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   130
val extract_info = fn thy => 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   131
 let val case_congs = map (#case_cong o #2) (!datatypes)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   132
         val case_rewrites = flat(map (#case_rewrites o #2) (!datatypes))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   133
 in {case_congs = #case_cong (#2 prod_record)::
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   134
                  #case_cong (#2 nat_record)::case_congs,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   135
     case_rewrites = #case_rewrites(#2 prod_record)@
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   136
                     #case_rewrites(#2 nat_record)@case_rewrites}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   137
 end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   138
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   139
end; (* Thry *)