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