TFL/rules.sml
changeset 6498 1ebbe18fe236
child 7262 a05dc63ca29b
equal deleted inserted replaced
6497:120ca2bb27e1 6498:1ebbe18fe236
       
     1 (*  Title:      TFL/rules
       
     2     ID:         $Id$
       
     3     Author:     Konrad Slind, Cambridge University Computer Laboratory
       
     4     Copyright   1997  University of Cambridge
       
     5 
       
     6 Emulation of HOL inference rules for TFL
       
     7 *)
       
     8 
       
     9 
       
    10 structure Rules : Rules_sig = 
       
    11 struct
       
    12 
       
    13 open Utils;
       
    14 
       
    15 structure USyntax  = USyntax;
       
    16 structure S  = USyntax;
       
    17 structure U  = Utils;
       
    18 structure D = Dcterm;
       
    19 
       
    20 
       
    21 fun RULES_ERR{func,mesg} = Utils.ERR{module = "Rules",func=func,mesg=mesg};
       
    22 
       
    23 
       
    24 fun cconcl thm = D.drop_prop(#prop(crep_thm thm));
       
    25 fun chyps thm = map D.drop_prop(#hyps(crep_thm thm));
       
    26 
       
    27 fun dest_thm thm = 
       
    28    let val {prop,hyps,...} = rep_thm thm
       
    29    in (map HOLogic.dest_Trueprop hyps, HOLogic.dest_Trueprop prop)
       
    30    end;
       
    31 
       
    32 
       
    33 
       
    34 (* Inference rules *)
       
    35 
       
    36 (*---------------------------------------------------------------------------
       
    37  *        Equality (one step)
       
    38  *---------------------------------------------------------------------------*)
       
    39 fun REFL tm = Thm.reflexive tm RS meta_eq_to_obj_eq;
       
    40 fun SYM thm = thm RS sym;
       
    41 
       
    42 fun ALPHA thm ctm1 =
       
    43    let val ctm2 = cprop_of thm
       
    44        val ctm2_eq = reflexive ctm2
       
    45        val ctm1_eq = reflexive ctm1
       
    46    in equal_elim (transitive ctm2_eq ctm1_eq) thm
       
    47    end;
       
    48 
       
    49 
       
    50 (*----------------------------------------------------------------------------
       
    51  *        typ instantiation
       
    52  *---------------------------------------------------------------------------*)
       
    53 fun INST_TYPE blist thm = 
       
    54   let val {sign,...} = rep_thm thm
       
    55       val blist' = map (fn (TVar(idx,_), B) => (idx, ctyp_of sign B)) blist
       
    56   in Thm.instantiate (blist',[]) thm
       
    57   end
       
    58   handle _ => raise RULES_ERR{func = "INST_TYPE", mesg = ""};
       
    59 
       
    60 
       
    61 (*----------------------------------------------------------------------------
       
    62  *        Implication and the assumption list
       
    63  *
       
    64  * Assumptions get stuck on the meta-language assumption list. Implications 
       
    65  * are in the object language, so discharging an assumption "A" from theorem
       
    66  * "B" results in something that looks like "A --> B".
       
    67  *---------------------------------------------------------------------------*)
       
    68 fun ASSUME ctm = Thm.assume (D.mk_prop ctm);
       
    69 
       
    70 
       
    71 (*---------------------------------------------------------------------------
       
    72  * Implication in TFL is -->. Meta-language implication (==>) is only used
       
    73  * in the implementation of some of the inference rules below.
       
    74  *---------------------------------------------------------------------------*)
       
    75 fun MP th1 th2 = th2 RS (th1 RS mp);
       
    76 
       
    77 (*forces the first argument to be a proposition if necessary*)
       
    78 fun DISCH tm thm = Thm.implies_intr (D.mk_prop tm) thm COMP impI;
       
    79 
       
    80 fun DISCH_ALL thm = Utils.itlist DISCH (#hyps (crep_thm thm)) thm;
       
    81 
       
    82 
       
    83 fun FILTER_DISCH_ALL P thm =
       
    84  let fun check tm = U.holds P (#t(rep_cterm tm))
       
    85  in  foldr (fn (tm,th) => if (check tm) then DISCH tm th else th)
       
    86               (chyps thm, thm)
       
    87  end;
       
    88 
       
    89 (* freezeT expensive! *)
       
    90 fun UNDISCH thm = 
       
    91    let val tm = D.mk_prop(#1(D.dest_imp(cconcl (freezeT thm))))
       
    92    in implies_elim (thm RS mp) (ASSUME tm)
       
    93    end
       
    94    handle _ => raise RULES_ERR{func = "UNDISCH", mesg = ""};
       
    95 
       
    96 fun PROVE_HYP ath bth =  MP (DISCH (cconcl ath) bth) ath;
       
    97 
       
    98 local val [p1,p2] = goal HOL.thy "(A-->B) ==> (B --> C) ==> (A-->C)"
       
    99       val dummy = by (rtac impI 1)
       
   100       val dummy = by (rtac (p2 RS mp) 1)
       
   101       val dummy = by (rtac (p1 RS mp) 1)
       
   102       val dummy = by (assume_tac 1)
       
   103       val imp_trans = result()
       
   104 in
       
   105 fun IMP_TRANS th1 th2 = th2 RS (th1 RS imp_trans)
       
   106 end;
       
   107 
       
   108 (*----------------------------------------------------------------------------
       
   109  *        Conjunction
       
   110  *---------------------------------------------------------------------------*)
       
   111 fun CONJUNCT1 thm = (thm RS conjunct1)
       
   112 fun CONJUNCT2 thm = (thm RS conjunct2);
       
   113 fun CONJUNCTS th  = (CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th))
       
   114                     handle _ => [th];
       
   115 
       
   116 fun LIST_CONJ [] = raise RULES_ERR{func = "LIST_CONJ", mesg = "empty list"}
       
   117   | LIST_CONJ [th] = th
       
   118   | LIST_CONJ (th::rst) = MP(MP(conjI COMP (impI RS impI)) th) (LIST_CONJ rst);
       
   119 
       
   120 
       
   121 (*----------------------------------------------------------------------------
       
   122  *        Disjunction
       
   123  *---------------------------------------------------------------------------*)
       
   124 local val {prop,sign,...} = rep_thm disjI1
       
   125       val [P,Q] = term_vars prop
       
   126       val disj1 = forall_intr (cterm_of sign Q) disjI1
       
   127 in
       
   128 fun DISJ1 thm tm = thm RS (forall_elim (D.drop_prop tm) disj1)
       
   129 end;
       
   130 
       
   131 local val {prop,sign,...} = rep_thm disjI2
       
   132       val [P,Q] = term_vars prop
       
   133       val disj2 = forall_intr (cterm_of sign P) disjI2
       
   134 in
       
   135 fun DISJ2 tm thm = thm RS (forall_elim (D.drop_prop tm) disj2)
       
   136 end;
       
   137 
       
   138 
       
   139 (*----------------------------------------------------------------------------
       
   140  *
       
   141  *                   A1 |- M1, ..., An |- Mn
       
   142  *     ---------------------------------------------------
       
   143  *     [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn]
       
   144  *
       
   145  *---------------------------------------------------------------------------*)
       
   146 
       
   147 
       
   148 fun EVEN_ORS thms =
       
   149   let fun blue ldisjs [] _ = []
       
   150         | blue ldisjs (th::rst) rdisjs =
       
   151             let val tail = tl rdisjs
       
   152                 val rdisj_tl = D.list_mk_disj tail
       
   153             in itlist DISJ2 ldisjs (DISJ1 th rdisj_tl)
       
   154                :: blue (ldisjs@[cconcl th]) rst tail
       
   155             end handle _ => [itlist DISJ2 ldisjs th]
       
   156    in
       
   157    blue [] thms (map cconcl thms)
       
   158    end;
       
   159 
       
   160 
       
   161 (*----------------------------------------------------------------------------
       
   162  *
       
   163  *         A |- P \/ Q   B,P |- R    C,Q |- R
       
   164  *     ---------------------------------------------------
       
   165  *                     A U B U C |- R
       
   166  *
       
   167  *---------------------------------------------------------------------------*)
       
   168 local val [p1,p2,p3] = goal HOL.thy "(P | Q) ==> (P --> R) ==> (Q --> R) ==> R"
       
   169       val dummy = by (rtac (p1 RS disjE) 1)
       
   170       val dummy = by (rtac (p2 RS mp) 1)
       
   171       val dummy = by (assume_tac 1)
       
   172       val dummy = by (rtac (p3 RS mp) 1)
       
   173       val dummy = by (assume_tac 1)
       
   174       val tfl_exE = result()
       
   175 in
       
   176 fun DISJ_CASES th1 th2 th3 = 
       
   177    let val c = D.drop_prop(cconcl th1)
       
   178        val (disj1,disj2) = D.dest_disj c
       
   179        val th2' = DISCH disj1 th2
       
   180        val th3' = DISCH disj2 th3
       
   181    in
       
   182    th3' RS (th2' RS (th1 RS tfl_exE))
       
   183    end
       
   184 end;
       
   185 
       
   186 
       
   187 (*-----------------------------------------------------------------------------
       
   188  *
       
   189  *       |- A1 \/ ... \/ An     [A1 |- M, ..., An |- M]
       
   190  *     ---------------------------------------------------
       
   191  *                           |- M
       
   192  *
       
   193  * Note. The list of theorems may be all jumbled up, so we have to 
       
   194  * first organize it to align with the first argument (the disjunctive 
       
   195  * theorem).
       
   196  *---------------------------------------------------------------------------*)
       
   197 
       
   198 fun organize eq =    (* a bit slow - analogous to insertion sort *)
       
   199  let fun extract a alist =
       
   200      let fun ex (_,[]) = raise RULES_ERR{func = "organize",
       
   201                                          mesg = "not a permutation.1"}
       
   202            | ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t)
       
   203      in ex ([],alist)
       
   204      end
       
   205      fun place [] [] = []
       
   206        | place (a::rst) alist =
       
   207            let val (item,next) = extract a alist
       
   208            in item::place rst next
       
   209            end
       
   210        | place _ _ = raise RULES_ERR{func = "organize",
       
   211                                      mesg = "not a permutation.2"}
       
   212  in place
       
   213  end;
       
   214 (* freezeT expensive! *)
       
   215 fun DISJ_CASESL disjth thl =
       
   216    let val c = cconcl disjth
       
   217        fun eq th atm = exists (fn t => HOLogic.dest_Trueprop t 
       
   218 			               aconv term_of atm)
       
   219 	                      (#hyps(rep_thm th))
       
   220        val tml = D.strip_disj c
       
   221        fun DL th [] = raise RULES_ERR{func="DISJ_CASESL",mesg="no cases"}
       
   222          | DL th [th1] = PROVE_HYP th th1
       
   223          | DL th [th1,th2] = DISJ_CASES th th1 th2
       
   224          | DL th (th1::rst) = 
       
   225             let val tm = #2(D.dest_disj(D.drop_prop(cconcl th)))
       
   226              in DISJ_CASES th th1 (DL (ASSUME tm) rst) end
       
   227    in DL (freezeT disjth) (organize eq tml thl)
       
   228    end;
       
   229 
       
   230 
       
   231 (*----------------------------------------------------------------------------
       
   232  *        Universals
       
   233  *---------------------------------------------------------------------------*)
       
   234 local (* this is fragile *)
       
   235       val {prop,sign,...} = rep_thm spec
       
   236       val x = hd (tl (term_vars prop))
       
   237       val (TVar (indx,_)) = type_of x
       
   238       val gspec = forall_intr (cterm_of sign x) spec
       
   239 in
       
   240 fun SPEC tm thm = 
       
   241    let val {sign,T,...} = rep_cterm tm
       
   242        val gspec' = instantiate([(indx,ctyp_of sign T)],[]) gspec
       
   243    in 
       
   244       thm RS (forall_elim tm gspec')
       
   245    end
       
   246 end;
       
   247 
       
   248 fun SPEC_ALL thm = rev_itlist SPEC (#1(D.strip_forall(cconcl thm))) thm;
       
   249 
       
   250 val ISPEC = SPEC
       
   251 val ISPECL = rev_itlist ISPEC;
       
   252 
       
   253 (* Not optimized! Too complicated. *)
       
   254 local val {prop,sign,...} = rep_thm allI
       
   255       val [P] = add_term_vars (prop, [])
       
   256       fun cty_theta s = map (fn (i,ty) => (i, ctyp_of s ty))
       
   257       fun ctm_theta s = map (fn (i,tm2) => 
       
   258                              let val ctm2 = cterm_of s tm2
       
   259                              in (cterm_of s (Var(i,#T(rep_cterm ctm2))), ctm2)
       
   260                              end)
       
   261       fun certify s (ty_theta,tm_theta) = (cty_theta s ty_theta, 
       
   262                                            ctm_theta s tm_theta)
       
   263 in
       
   264 fun GEN v th =
       
   265    let val gth = forall_intr v th
       
   266        val {prop=Const("all",_)$Abs(x,ty,rst),sign,...} = rep_thm gth
       
   267        val P' = Abs(x,ty, HOLogic.dest_Trueprop rst)  (* get rid of trueprop *)
       
   268        val tsig = #tsig(Sign.rep_sg sign)
       
   269        val theta = Pattern.match tsig (P,P')
       
   270        val allI2 = instantiate (certify sign theta) allI
       
   271        val thm = implies_elim allI2 gth
       
   272        val {prop = tp $ (A $ Abs(_,_,M)),sign,...} = rep_thm thm
       
   273        val prop' = tp $ (A $ Abs(x,ty,M))
       
   274    in ALPHA thm (cterm_of sign prop')
       
   275    end
       
   276 end;
       
   277 
       
   278 val GENL = itlist GEN;
       
   279 
       
   280 fun GEN_ALL thm = 
       
   281    let val {prop,sign,...} = rep_thm thm
       
   282        val tycheck = cterm_of sign
       
   283        val vlist = map tycheck (add_term_vars (prop, []))
       
   284   in GENL vlist thm
       
   285   end;
       
   286 
       
   287 
       
   288 fun MATCH_MP th1 th2 = 
       
   289    if (D.is_forall (D.drop_prop(cconcl th1)))
       
   290    then MATCH_MP (th1 RS spec) th2
       
   291    else MP th1 th2;
       
   292 
       
   293 
       
   294 (*----------------------------------------------------------------------------
       
   295  *        Existentials
       
   296  *---------------------------------------------------------------------------*)
       
   297 
       
   298 
       
   299 
       
   300 (*--------------------------------------------------------------------------- 
       
   301  * Existential elimination
       
   302  *
       
   303  *      A1 |- ?x.t[x]   ,   A2, "t[v]" |- t'
       
   304  *      ------------------------------------     (variable v occurs nowhere)
       
   305  *                A1 u A2 |- t'
       
   306  *
       
   307  *---------------------------------------------------------------------------*)
       
   308 
       
   309 local val [p1,p2] = goal HOL.thy "(? x. P x) ==> (!x. P x --> Q) ==> Q"
       
   310       val dummy = by (rtac (p1 RS exE) 1)
       
   311       val dummy = by (rtac ((p2 RS allE) RS mp) 1)
       
   312       val dummy = by (assume_tac 2)
       
   313       val dummy = by (assume_tac 1)
       
   314       val choose_thm = result()
       
   315 in
       
   316 fun CHOOSE(fvar,exth) fact =
       
   317    let val lam = #2(dest_comb(D.drop_prop(cconcl exth)))
       
   318        val redex = capply lam fvar
       
   319        val {sign, t = t$u,...} = rep_cterm redex
       
   320        val residue = cterm_of sign (betapply(t,u))
       
   321     in GEN fvar (DISCH residue fact)  RS (exth RS choose_thm)
       
   322    end
       
   323 end;
       
   324 
       
   325 
       
   326 local val {prop,sign,...} = rep_thm exI
       
   327       val [P,x] = term_vars prop
       
   328 in
       
   329 fun EXISTS (template,witness) thm =
       
   330    let val {prop,sign,...} = rep_thm thm
       
   331        val P' = cterm_of sign P
       
   332        val x' = cterm_of sign x
       
   333        val abstr = #2(dest_comb template)
       
   334    in
       
   335    thm RS (cterm_instantiate[(P',abstr), (x',witness)] exI)
       
   336    end
       
   337 end;
       
   338 
       
   339 (*----------------------------------------------------------------------------
       
   340  *
       
   341  *         A |- M
       
   342  *   -------------------   [v_1,...,v_n]
       
   343  *    A |- ?v1...v_n. M
       
   344  *
       
   345  *---------------------------------------------------------------------------*)
       
   346 
       
   347 fun EXISTL vlist th = 
       
   348   U.itlist (fn v => fn thm => EXISTS(D.mk_exists(v,cconcl thm), v) thm)
       
   349            vlist th;
       
   350 
       
   351 
       
   352 (*----------------------------------------------------------------------------
       
   353  *
       
   354  *       A |- M[x_1,...,x_n]
       
   355  *   ----------------------------   [(x |-> y)_1,...,(x |-> y)_n]
       
   356  *       A |- ?y_1...y_n. M
       
   357  *
       
   358  *---------------------------------------------------------------------------*)
       
   359 (* Could be improved, but needs "subst_free" for certified terms *)
       
   360 
       
   361 fun IT_EXISTS blist th = 
       
   362    let val {sign,...} = rep_thm th
       
   363        val tych = cterm_of sign
       
   364        val detype = #t o rep_cterm
       
   365        val blist' = map (fn (x,y) => (detype x, detype y)) blist
       
   366        fun ?v M  = cterm_of sign (S.mk_exists{Bvar=v,Body = M})
       
   367        
       
   368   in
       
   369   U.itlist (fn (b as (r1,r2)) => fn thm => 
       
   370         EXISTS(?r2(subst_free[b] 
       
   371 		   (HOLogic.dest_Trueprop(#prop(rep_thm thm)))), tych r1)
       
   372               thm)
       
   373        blist' th
       
   374   end;
       
   375 
       
   376 (*---------------------------------------------------------------------------
       
   377  *  Faster version, that fails for some as yet unknown reason
       
   378  * fun IT_EXISTS blist th = 
       
   379  *    let val {sign,...} = rep_thm th
       
   380  *        val tych = cterm_of sign
       
   381  *        fun detype (x,y) = ((#t o rep_cterm) x, (#t o rep_cterm) y)
       
   382  *   in
       
   383  *  fold (fn (b as (r1,r2), thm) => 
       
   384  *  EXISTS(D.mk_exists(r2, tych(subst_free[detype b](#t(rep_cterm(cconcl thm))))),
       
   385  *           r1) thm)  blist th
       
   386  *   end;
       
   387  *---------------------------------------------------------------------------*)
       
   388 
       
   389 (*----------------------------------------------------------------------------
       
   390  *        Rewriting
       
   391  *---------------------------------------------------------------------------*)
       
   392 
       
   393 fun SUBS thl = 
       
   394    rewrite_rule (map (fn th => (th RS eq_reflection) handle _ => th) thl);
       
   395 
       
   396 local fun rew_conv mss = Thm.rewrite_cterm (true,false,false) mss (K(K None))
       
   397 in
       
   398 fun simpl_conv ss thl ctm = 
       
   399  rew_conv (Thm.mss_of (#simps (Thm.dest_mss (#mss (rep_ss ss))) @ thl)) ctm
       
   400  RS meta_eq_to_obj_eq
       
   401 end;
       
   402 
       
   403 local fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1])
       
   404 in
       
   405 val RIGHT_ASSOC = rewrite_rule [prover"((a|b)|c) = (a|(b|c))" RS eq_reflection]
       
   406 val ASM = refl RS iffD1
       
   407 end;
       
   408 
       
   409 
       
   410 
       
   411 
       
   412 (*---------------------------------------------------------------------------
       
   413  *                  TERMINATION CONDITION EXTRACTION
       
   414  *---------------------------------------------------------------------------*)
       
   415 
       
   416 
       
   417 (* Object language quantifier, i.e., "!" *)
       
   418 fun Forall v M = S.mk_forall{Bvar=v, Body=M};
       
   419 
       
   420 
       
   421 (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *)
       
   422 fun is_cong thm = 
       
   423   let val {prop, ...} = rep_thm thm
       
   424   in case prop 
       
   425      of (Const("==>",_)$(Const("Trueprop",_)$ _) $
       
   426          (Const("==",_) $ (Const ("cut",_) $ f $ R $ a $ x) $ _)) => false
       
   427       | _ => true
       
   428   end;
       
   429 
       
   430 
       
   431    
       
   432 fun dest_equal(Const ("==",_) $ 
       
   433 	       (Const ("Trueprop",_) $ lhs) 
       
   434 	       $ (Const ("Trueprop",_) $ rhs)) = {lhs=lhs, rhs=rhs}
       
   435   | dest_equal(Const ("==",_) $ lhs $ rhs)  = {lhs=lhs, rhs=rhs}
       
   436   | dest_equal tm = S.dest_eq tm;
       
   437 
       
   438 fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm));
       
   439 
       
   440 fun dest_all(Const("all",_) $ (a as Abs _)) = S.dest_abs a
       
   441   | dest_all _ = raise RULES_ERR{func = "dest_all", mesg = "not a !!"};
       
   442 
       
   443 val is_all = Utils.can dest_all;
       
   444 
       
   445 fun strip_all fm =
       
   446    if (is_all fm)
       
   447    then let val {Bvar,Body} = dest_all fm
       
   448             val (bvs,core)  = strip_all Body
       
   449         in ((Bvar::bvs), core)
       
   450         end
       
   451    else ([],fm);
       
   452 
       
   453 fun break_all(Const("all",_) $ Abs (_,_,body)) = body
       
   454   | break_all _ = raise RULES_ERR{func = "break_all", mesg = "not a !!"};
       
   455 
       
   456 fun list_break_all(Const("all",_) $ Abs (s,ty,body)) = 
       
   457      let val (L,core) = list_break_all body
       
   458      in ((s,ty)::L, core)
       
   459      end
       
   460   | list_break_all tm = ([],tm);
       
   461 
       
   462 (*---------------------------------------------------------------------------
       
   463  * Rename a term of the form
       
   464  *
       
   465  *      !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn 
       
   466  *                  ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn.
       
   467  * to one of
       
   468  *
       
   469  *      !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn 
       
   470  *      ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn.
       
   471  * 
       
   472  * This prevents name problems in extraction, and helps the result to read
       
   473  * better. There is a problem with varstructs, since they can introduce more
       
   474  * than n variables, and some extra reasoning needs to be done.
       
   475  *---------------------------------------------------------------------------*)
       
   476 
       
   477 fun get ([],_,L) = rev L
       
   478   | get (ant::rst,n,L) =  
       
   479       case (list_break_all ant)
       
   480         of ([],_) => get (rst, n+1,L)
       
   481          | (vlist,body) =>
       
   482             let val eq = Logic.strip_imp_concl body
       
   483                 val (f,args) = S.strip_comb (get_lhs eq)
       
   484                 val (vstrl,_) = S.strip_abs f
       
   485                 val names  = variantlist (map (#1 o dest_Free) vstrl,
       
   486 					  add_term_names(body, []))
       
   487             in get (rst, n+1, (names,n)::L)
       
   488             end handle _ => get (rst, n+1, L);
       
   489 
       
   490 (* Note: rename_params_rule counts from 1, not 0 *)
       
   491 fun rename thm = 
       
   492   let val {prop,sign,...} = rep_thm thm
       
   493       val tych = cterm_of sign
       
   494       val ants = Logic.strip_imp_prems prop
       
   495       val news = get (ants,1,[])
       
   496   in 
       
   497   U.rev_itlist rename_params_rule news thm
       
   498   end;
       
   499 
       
   500 
       
   501 (*---------------------------------------------------------------------------
       
   502  * Beta-conversion to the rhs of an equation (taken from hol90/drule.sml)
       
   503  *---------------------------------------------------------------------------*)
       
   504 
       
   505 fun list_beta_conv tm =
       
   506   let fun rbeta th = transitive th (beta_conversion(#2(D.dest_eq(cconcl th))))
       
   507       fun iter [] = reflexive tm
       
   508         | iter (v::rst) = rbeta (combination(iter rst) (reflexive v))
       
   509   in iter  end;
       
   510 
       
   511 
       
   512 (*---------------------------------------------------------------------------
       
   513  * Trace information for the rewriter
       
   514  *---------------------------------------------------------------------------*)
       
   515 val term_ref = ref[] : term list ref
       
   516 val mss_ref = ref [] : meta_simpset list ref;
       
   517 val thm_ref = ref [] : thm list ref;
       
   518 val tracing = ref false;
       
   519 
       
   520 fun say s = if !tracing then writeln s else ();
       
   521 
       
   522 fun print_thms s L = 
       
   523   say (cat_lines (s :: map string_of_thm L));
       
   524 
       
   525 fun print_cterms s L = 
       
   526   say (cat_lines (s :: map string_of_cterm L));
       
   527 
       
   528 
       
   529 (*---------------------------------------------------------------------------
       
   530  * General abstraction handlers, should probably go in USyntax.
       
   531  *---------------------------------------------------------------------------*)
       
   532 fun mk_aabs(vstr,body) = S.mk_abs{Bvar=vstr,Body=body}
       
   533                          handle _ => S.mk_pabs{varstruct = vstr, body = body};
       
   534 
       
   535 fun list_mk_aabs (vstrl,tm) =
       
   536     U.itlist (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm;
       
   537 
       
   538 fun dest_aabs tm = 
       
   539    let val {Bvar,Body} = S.dest_abs tm
       
   540    in (Bvar,Body)
       
   541    end handle _ => let val {varstruct,body} = S.dest_pabs tm
       
   542                    in (varstruct,body)
       
   543                    end;
       
   544 
       
   545 fun strip_aabs tm =
       
   546    let val (vstr,body) = dest_aabs tm
       
   547        val (bvs, core) = strip_aabs body
       
   548    in (vstr::bvs, core)
       
   549    end
       
   550    handle _ => ([],tm);
       
   551 
       
   552 fun dest_combn tm 0 = (tm,[])
       
   553   | dest_combn tm n = 
       
   554      let val {Rator,Rand} = S.dest_comb tm
       
   555          val (f,rands) = dest_combn Rator (n-1)
       
   556      in (f,Rand::rands)
       
   557      end;
       
   558 
       
   559 
       
   560 
       
   561 
       
   562 local fun dest_pair M = let val {fst,snd} = S.dest_pair M in (fst,snd) end
       
   563       fun mk_fst tm = 
       
   564           let val ty as Type("*", [fty,sty]) = type_of tm
       
   565           in  Const ("fst", ty --> fty) $ tm  end
       
   566       fun mk_snd tm = 
       
   567           let val ty as Type("*", [fty,sty]) = type_of tm
       
   568           in  Const ("snd", ty --> sty) $ tm  end
       
   569 in
       
   570 fun XFILL tych x vstruct = 
       
   571   let fun traverse p xocc L =
       
   572         if (is_Free p)
       
   573         then tych xocc::L
       
   574         else let val (p1,p2) = dest_pair p
       
   575              in traverse p1 (mk_fst xocc) (traverse p2  (mk_snd xocc) L)
       
   576              end
       
   577   in 
       
   578   traverse vstruct x []
       
   579 end end;
       
   580 
       
   581 (*---------------------------------------------------------------------------
       
   582  * Replace a free tuple (vstr) by a universally quantified variable (a).
       
   583  * Note that the notion of "freeness" for a tuple is different than for a
       
   584  * variable: if variables in the tuple also occur in any other place than
       
   585  * an occurrences of the tuple, they aren't "free" (which is thus probably
       
   586  *  the wrong word to use).
       
   587  *---------------------------------------------------------------------------*)
       
   588 
       
   589 fun VSTRUCT_ELIM tych a vstr th = 
       
   590   let val L = S.free_vars_lr vstr
       
   591       val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr)))
       
   592       val thm1 = implies_intr bind1 (SUBS [SYM(assume bind1)] th)
       
   593       val thm2 = forall_intr_list (map tych L) thm1
       
   594       val thm3 = forall_elim_list (XFILL tych a vstr) thm2
       
   595   in refl RS
       
   596      rewrite_rule[symmetric (surjective_pairing RS eq_reflection)] thm3
       
   597   end;
       
   598 
       
   599 fun PGEN tych a vstr th = 
       
   600   let val a1 = tych a
       
   601       val vstr1 = tych vstr
       
   602   in
       
   603   forall_intr a1 
       
   604      (if (is_Free vstr) 
       
   605       then cterm_instantiate [(vstr1,a1)] th
       
   606       else VSTRUCT_ELIM tych a vstr th)
       
   607   end;
       
   608 
       
   609 
       
   610 (*---------------------------------------------------------------------------
       
   611  * Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into
       
   612  *
       
   613  *     (([x,y],N),vstr)
       
   614  *---------------------------------------------------------------------------*)
       
   615 fun dest_pbeta_redex M n = 
       
   616   let val (f,args) = dest_combn M n
       
   617       val dummy = dest_aabs f
       
   618   in (strip_aabs f,args)
       
   619   end;
       
   620 
       
   621 fun pbeta_redex M n = U.can (U.C dest_pbeta_redex n) M;
       
   622 
       
   623 fun dest_impl tm = 
       
   624   let val ants = Logic.strip_imp_prems tm
       
   625       val eq = Logic.strip_imp_concl tm
       
   626   in (ants,get_lhs eq)
       
   627   end;
       
   628 
       
   629 fun restricted t = is_some (S.find_term
       
   630 			    (fn (Const("cut",_)) =>true | _ => false) 
       
   631 			    t)
       
   632 
       
   633 fun CONTEXT_REWRITE_RULE (func, G, cut_lemma, congs) th =
       
   634  let val globals = func::G
       
   635      val pbeta_reduce = simpl_conv empty_ss [split RS eq_reflection];
       
   636      val tc_list = ref[]: term list ref
       
   637      val dummy = term_ref := []
       
   638      val dummy = thm_ref  := []
       
   639      val dummy = mss_ref  := []
       
   640      val cut_lemma' = cut_lemma RS eq_reflection
       
   641      fun prover mss thm =
       
   642      let fun cong_prover mss thm =
       
   643          let val dummy = say "cong_prover:"
       
   644              val cntxt = prems_of_mss mss
       
   645              val dummy = print_thms "cntxt:" cntxt
       
   646              val dummy = say "cong rule:"
       
   647              val dummy = say (string_of_thm thm)
       
   648              val dummy = thm_ref := (thm :: !thm_ref)
       
   649              val dummy = mss_ref := (mss :: !mss_ref)
       
   650              (* Unquantified eliminate *)
       
   651              fun uq_eliminate (thm,imp,sign) = 
       
   652                  let val tych = cterm_of sign
       
   653                      val dummy = print_cterms "To eliminate:" [tych imp]
       
   654                      val ants = map tych (Logic.strip_imp_prems imp)
       
   655                      val eq = Logic.strip_imp_concl imp
       
   656                      val lhs = tych(get_lhs eq)
       
   657                      val mss' = add_prems(mss, map ASSUME ants)
       
   658                      val lhs_eq_lhs1 = Thm.rewrite_cterm(false,true,false)mss' prover lhs
       
   659                        handle _ => reflexive lhs
       
   660                      val dummy = print_thms "proven:" [lhs_eq_lhs1]
       
   661                      val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1
       
   662                      val lhs_eeq_lhs2 = lhs_eq_lhs2 RS meta_eq_to_obj_eq
       
   663                   in
       
   664                   lhs_eeq_lhs2 COMP thm
       
   665                   end
       
   666              fun pq_eliminate (thm,sign,vlist,imp_body,lhs_eq) =
       
   667               let val ((vstrl,_),args) = dest_pbeta_redex lhs_eq(length vlist)
       
   668                   val dummy = assert (forall (op aconv)
       
   669                                       (ListPair.zip (vlist, args)))
       
   670                                "assertion failed in CONTEXT_REWRITE_RULE"
       
   671                   val imp_body1 = subst_free (ListPair.zip (args, vstrl))
       
   672                                              imp_body
       
   673                   val tych = cterm_of sign
       
   674                   val ants1 = map tych (Logic.strip_imp_prems imp_body1)
       
   675                   val eq1 = Logic.strip_imp_concl imp_body1
       
   676                   val Q = get_lhs eq1
       
   677                   val QeqQ1 = pbeta_reduce (tych Q)
       
   678                   val Q1 = #2(D.dest_eq(cconcl QeqQ1))
       
   679                   val mss' = add_prems(mss, map ASSUME ants1)
       
   680                   val Q1eeqQ2 = Thm.rewrite_cterm (false,true,false) mss' prover Q1
       
   681                                 handle _ => reflexive Q1
       
   682                   val Q2 = #2 (Logic.dest_equals (#prop(rep_thm Q1eeqQ2)))
       
   683                   val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl))
       
   684                   val Q2eeqQ3 = symmetric(pbeta_reduce Q3 RS eq_reflection)
       
   685                   val thA = transitive(QeqQ1 RS eq_reflection) Q1eeqQ2
       
   686                   val QeeqQ3 = transitive thA Q2eeqQ3 handle _ =>
       
   687                                ((Q2eeqQ3 RS meta_eq_to_obj_eq) 
       
   688                                 RS ((thA RS meta_eq_to_obj_eq) RS trans))
       
   689                                 RS eq_reflection
       
   690                   val impth = implies_intr_list ants1 QeeqQ3
       
   691                   val impth1 = impth RS meta_eq_to_obj_eq
       
   692                   (* Need to abstract *)
       
   693                   val ant_th = U.itlist2 (PGEN tych) args vstrl impth1
       
   694               in ant_th COMP thm
       
   695               end
       
   696              fun q_eliminate (thm,imp,sign) =
       
   697               let val (vlist,imp_body) = strip_all imp
       
   698                   val (ants,Q) = dest_impl imp_body
       
   699               in if (pbeta_redex Q) (length vlist)
       
   700                  then pq_eliminate (thm,sign,vlist,imp_body,Q)
       
   701                  else 
       
   702                  let val tych = cterm_of sign
       
   703                      val ants1 = map tych ants
       
   704                      val mss' = add_prems(mss, map ASSUME ants1)
       
   705                      val Q_eeq_Q1 = Thm.rewrite_cterm(false,true,false) mss' 
       
   706                                                      prover (tych Q)
       
   707                       handle _ => reflexive (tych Q)
       
   708                      val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1
       
   709                      val lhs_eq_lhs2 = lhs_eeq_lhs2 RS meta_eq_to_obj_eq
       
   710                      val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2
       
   711                  in
       
   712                  ant_th COMP thm
       
   713               end end
       
   714 
       
   715              fun eliminate thm = 
       
   716                case (rep_thm thm)
       
   717                of {prop = (Const("==>",_) $ imp $ _), sign, ...} =>
       
   718                    eliminate
       
   719                     (if not(is_all imp)
       
   720                      then uq_eliminate (thm,imp,sign)
       
   721                      else q_eliminate (thm,imp,sign))
       
   722                             (* Assume that the leading constant is ==,   *)
       
   723                 | _ => thm  (* if it is not a ==>                        *)
       
   724          in Some(eliminate (rename thm))
       
   725          end handle _ => None
       
   726 
       
   727         fun restrict_prover mss thm =
       
   728           let val dummy = say "restrict_prover:"
       
   729               val cntxt = rev(prems_of_mss mss)
       
   730               val dummy = print_thms "cntxt:" cntxt
       
   731               val {prop = Const("==>",_) $ (Const("Trueprop",_) $ A) $ _,
       
   732                    sign,...} = rep_thm thm
       
   733               fun genl tm = let val vlist = gen_rems (op aconv)
       
   734                                            (add_term_frees(tm,[]), globals)
       
   735                             in U.itlist Forall vlist tm
       
   736                             end
       
   737               (*--------------------------------------------------------------
       
   738                * This actually isn't quite right, since it will think that
       
   739                * not-fully applied occs. of "f" in the context mean that the
       
   740                * current call is nested. The real solution is to pass in a
       
   741                * term "f v1..vn" which is a pattern that any full application
       
   742                * of "f" will match.
       
   743                *-------------------------------------------------------------*)
       
   744               val func_name = #1(dest_Const func)
       
   745               fun is_func (Const (name,_)) = (name = func_name)
       
   746 		| is_func _                = false
       
   747               val rcontext = rev cntxt
       
   748               val cncl = HOLogic.dest_Trueprop o #prop o rep_thm
       
   749               val antl = case rcontext of [] => [] 
       
   750                          | _   => [S.list_mk_conj(map cncl rcontext)]
       
   751               val TC = genl(S.list_mk_imp(antl, A))
       
   752               val dummy = print_cterms "func:" [cterm_of sign func]
       
   753               val dummy = print_cterms "TC:" 
       
   754 		              [cterm_of sign (HOLogic.mk_Trueprop TC)]
       
   755               val dummy = tc_list := (TC :: !tc_list)
       
   756               val nestedp = is_some (S.find_term is_func TC)
       
   757               val dummy = if nestedp then say "nested" else say "not_nested"
       
   758               val dummy = term_ref := ([func,TC]@(!term_ref))
       
   759               val th' = if nestedp then raise RULES_ERR{func = "solver", 
       
   760                                                       mesg = "nested function"}
       
   761                         else let val cTC = cterm_of sign 
       
   762 			                      (HOLogic.mk_Trueprop TC)
       
   763                              in case rcontext of
       
   764                                 [] => SPEC_ALL(ASSUME cTC)
       
   765                                | _ => MP (SPEC_ALL (ASSUME cTC)) 
       
   766                                          (LIST_CONJ rcontext)
       
   767                              end
       
   768               val th'' = th' RS thm
       
   769           in Some (th'')
       
   770           end handle _ => None
       
   771     in
       
   772     (if (is_cong thm) then cong_prover else restrict_prover) mss thm
       
   773     end
       
   774     val ctm = cprop_of th
       
   775     val th1 = Thm.rewrite_cterm(false,true,false) (add_congs(mss_of [cut_lemma'], congs))
       
   776                             prover ctm
       
   777     val th2 = equal_elim th1 th
       
   778  in
       
   779  (th2, filter (not o restricted) (!tc_list))
       
   780  end;
       
   781 
       
   782 
       
   783 
       
   784 fun prove (ptm,tac) = 
       
   785     #1 (freeze_thaw (prove_goalw_cterm [] ptm (fn _ => [tac])));
       
   786 
       
   787 
       
   788 end; (* Rules *)