src/HOL/add_ind_def.ML
author nipkow
Mon, 03 Nov 1997 09:58:06 +0100
changeset 4072 d0d32dd77440
parent 4027 15768dba480e
child 4235 c4f1d3940d95
permissions -rw-r--r--
expand_option_bind -> split_option_bind
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
     1
(*  Title:      HOL/add_ind_def.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Fixedpoint definition module -- for Inductive/Coinductive Definitions
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
Features:
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
* least or greatest fixedpoints
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
* user-specified product and sum constructions
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
* mutually recursive definitions
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
* definitions involving arbitrary monotone operators
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
* automatically proves introduction and elimination rules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
The recursive sets must *already* be declared as constants in parent theory!
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
  Introduction rules have the form
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
  [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
  where M is some monotone operator (usually the identity)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
  P(x) is any (non-conjunctive) side condition on the free variables
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
  ti, t are any terms
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
  Sj, Sk are two of the sets being defined in mutual recursion
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
Sums are used only for mutual recursion;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
Products are used only to derive "streamlined" induction rules for relations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
Nestings of disjoint sum types:
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
   (a+(b+c)) for 3,  ((a+b)+(c+d)) for 4,  ((a+b)+(c+(d+e))) for 5,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
   ((a+(b+c))+(d+(e+f))) for 6
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    32
signature FP =          (** Description of a fixed point operator **)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
  sig
2859
7d640451ae7d Now tests for essential ancestors (Lfp or Gfp)
paulson
parents: 2270
diff changeset
    34
  val checkThy  : theory -> unit   (*signals error if Lfp/Gfp is missing*)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    35
  val oper      : string * typ * term -> term   (*fixed point operator*)
2859
7d640451ae7d Now tests for essential ancestors (Lfp or Gfp)
paulson
parents: 2270
diff changeset
    36
  val Tarski    : thm              (*Tarski's fixed point theorem*)
7d640451ae7d Now tests for essential ancestors (Lfp or Gfp)
paulson
parents: 2270
diff changeset
    37
  val induct    : thm              (*induction/coinduction rule*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
signature ADD_INDUCTIVE_DEF =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
  sig 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
  val add_fp_def_i : term list * term list -> theory -> theory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
(*Declares functions to add fixedpoint/constructor defs to a theory*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
functor Add_inductive_def_Fun (Fp: FP) : ADD_INDUCTIVE_DEF =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
struct
1397
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
    51
open Ind_Syntax;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
(*internal version*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
fun add_fp_def_i (rec_tms, intr_tms) thy = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
  let
2859
7d640451ae7d Now tests for essential ancestors (Lfp or Gfp)
paulson
parents: 2270
diff changeset
    56
    val dummy = Fp.checkThy thy		(*has essential ancestors?*)
7d640451ae7d Now tests for essential ancestors (Lfp or Gfp)
paulson
parents: 2270
diff changeset
    57
    
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
    val sign = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
1189
c17a8fd0c95d Changed comments
lcp
parents: 923
diff changeset
    60
    (*rec_params should agree for all mutually recursive components*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
    val rec_hds = map head_of rec_tms;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
    val _ = assert_all is_Const rec_hds
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    64
            (fn t => "Recursive set not previously declared as constant: " ^ 
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    65
                     Sign.string_of_term sign t);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
    (*Now we know they are all Consts, so get their names, type and params*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
    val rec_names = map (#1 o dest_Const) rec_hds
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
    and (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
3945
ae9c61d69888 adapted to qualified names;
wenzelm
parents: 3768
diff changeset
    71
    val rec_base_names = map Sign.base_name rec_names;
ae9c61d69888 adapted to qualified names;
wenzelm
parents: 3768
diff changeset
    72
    val _ = assert_all Syntax.is_identifier rec_base_names
ae9c61d69888 adapted to qualified names;
wenzelm
parents: 3768
diff changeset
    73
       (fn a => "Base name of recursive set not an identifier: " ^ a);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
    local (*Checking the introduction rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
      val intr_sets = map (#2 o rule_concl_msg sign) intr_tms;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
      fun intr_ok set =
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    78
          case head_of set of Const(a,_) => a mem rec_names | _ => false;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
    in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
      val _ =  assert_all intr_ok intr_sets
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    81
         (fn t => "Conclusion of rule does not name a recursive set: " ^ 
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    82
                  Sign.string_of_term sign t);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
    end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
    val _ = assert_all is_Free rec_params
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    86
        (fn t => "Param in recursion term not a free variable: " ^
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
    87
                 Sign.string_of_term sign t);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
    (*** Construct the lfp definition ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
    val mk_variant = variant (foldr add_term_names (intr_tms,[]));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
    val z = mk_variant"z" and X = mk_variant"X" and w = mk_variant"w";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
1189
c17a8fd0c95d Changed comments
lcp
parents: 923
diff changeset
    94
    (*Mutual recursion ?? *)
2995
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
    95
    val dom_set  = body_type recT
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
    96
    val dom_sumT = dest_setT dom_set
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
    val freez   = Free(z, dom_sumT)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
    and freeX   = Free(X, dom_set);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
    fun dest_tprop (Const("Trueprop",_) $ P) = P
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
      | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   103
                              Sign.string_of_term sign Q);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
    (*Makes a disjunct from an introduction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
    fun lfp_part intr = (*quantify over rule's free vars except parameters*)
1397
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
   107
      let val prems = map dest_tprop (Logic.strip_imp_prems intr)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   108
          val _ = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   109
          val exfrees = term_frees intr \\ rec_params
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   110
          val zeq = eq_const dom_sumT $ freez $ (#1 (rule_concl intr))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
      in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
    (*The Part(A,h) terms -- compose injections to make h*)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   114
    fun mk_Part (Bound 0, _) = freeX    (*no mutual rec, no Part needed*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
      | mk_Part (h, domT)    = 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   116
          let val goodh = mend_sum_types (h, dom_sumT)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
              and Part_const = 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   118
                  Const("Part", [dom_set, domT-->dom_sumT]---> dom_set)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
          in  Part_const $ freeX $ Abs(w,domT,goodh)  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
2995
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
   121
    (*Access to balanced disjoint sums via injections??
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
   122
      Mutual recursion is NOT supported*)
2270
d7513875b2b8 Replaced map...~~ by ListPair.map
paulson
parents: 1465
diff changeset
   123
    val parts = ListPair.map mk_Part
2995
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
   124
                (accesses_bal (ap Inl, ap Inr, Bound 0) 1,
84df3b150b67 Disabled the attempts for mutual induction to work so that single induction
paulson
parents: 2859
diff changeset
   125
                 [dom_set]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
    (*replace each set by the corresponding Part(A,h)*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
    val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
    val lfp_rhs = Fp.oper(X, dom_sumT, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   131
                          mk_Collect(z, dom_sumT, 
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   132
                                     fold_bal (app disj) part_intrs))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
    (*** Make the new theory ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
    (*A key definition:
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
      If no mutual recursion then it equals the one recursive set.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
      If mutual recursion then it differs from all the recursive sets. *)
3945
ae9c61d69888 adapted to qualified names;
wenzelm
parents: 3768
diff changeset
   140
    val big_rec_base_name = space_implode "_" rec_base_names;
ae9c61d69888 adapted to qualified names;
wenzelm
parents: 3768
diff changeset
   141
    val big_rec_name = Sign.full_name sign big_rec_base_name;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
    (*Big_rec... is the union of the mutually recursive sets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
    val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   146
    (*The individual sets must already be declared*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
    val axpairs = map mk_defpair 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   148
          ((big_rec_tm, lfp_rhs) ::
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   149
           (case parts of 
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   150
               [_] => []                        (*no mutual recursion*)
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   151
             | _ => rec_tms ~~          (*define the sets as Parts*)
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   152
                    map (subst_atomic [(freeX, big_rec_tm)]) parts));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
    val _ = seq (writeln o Sign.string_of_term sign o #2) axpairs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
  
1397
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
   156
    (*Detect occurrences of operator, even with other types!*)
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
   157
    val _ = (case rec_names inter (add_term_names (lfp_rhs,[])) of
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   158
               [] => ()
5d7a7e439cec expanded tabs
clasohm
parents: 1397
diff changeset
   159
             | x::_ => error ("Illegal occurrence of recursion op: " ^ x ^
1397
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
   160
                               "\n\t*Consider adding type constraints*"))
b010f04fcb9c Improved error message, suggesting addition of
paulson
parents: 1189
diff changeset
   161
4027
15768dba480e PureThy.add_store_defs_i;
wenzelm
parents: 3945
diff changeset
   162
  in  thy |> PureThy.add_store_defs_i axpairs  end
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
end;