src/FOLP/simp.ML
author wenzelm
Fri, 06 Mar 2015 15:58:56 +0100
changeset 59621 291934bac95e
parent 59582 0fbed69ff081
child 60644 4af8b9c2b52f
permissions -rw-r--r--
Thm.cterm_of and Thm.ctyp_of operate on local context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 26939
diff changeset
     1
(*  Title:      FOLP/simp.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    Author:     Tobias Nipkow
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
FOLP version of...
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
Generic simplifier, suitable for most logics.  (from Provers)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
This version allows instantiation of Vars in the subgoal, since the proof
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
term must change.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
signature SIMP_DATA =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
sig
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
  val case_splits  : (thm * string) list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
  val dest_red     : term -> term * term * term
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
  val mk_rew_rules : thm -> thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
  val norm_thms    : (thm*thm) list (* [(?x>>norm(?x), norm(?x)>>?x), ...] *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  val red1         : thm        (*  ?P>>?Q  ==>  ?P  ==>  ?Q  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
  val red2         : thm        (*  ?P>>?Q  ==>  ?Q  ==>  ?P  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
  val refl_thms    : thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
  val subst_thms   : thm list   (* [ ?a>>?b ==> ?P(?a) ==> ?P(?b), ...] *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
  val trans_thms   : thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
infix 4 addrews addcongs delrews delcongs setauto;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
signature SIMP =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
sig
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
  type simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
  val empty_ss  : simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
  val addcongs  : simpset * thm list -> simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
  val addrews   : simpset * thm list -> simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
  val delcongs  : simpset * thm list -> simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
  val delrews   : simpset * thm list -> simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
  val dest_ss   : simpset -> thm list * thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
  val print_ss  : simpset -> unit
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
  val setauto   : simpset * (int -> tactic) -> simpset
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
  val ASM_SIMP_CASE_TAC : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
  val ASM_SIMP_TAC      : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
  val CASE_TAC          : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
  val SIMP_CASE2_TAC    : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
  val SIMP_THM          : simpset -> thm -> thm
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
  val SIMP_TAC          : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
  val SIMP_CASE_TAC     : simpset -> int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
  val mk_congs          : theory -> string list -> thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
  val mk_typed_congs    : theory -> (string * string) list -> thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
(* temporarily disabled:
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
  val extract_free_congs        : unit -> thm list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
*)
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32449
diff changeset
    52
  val tracing   : bool Unsynchronized.ref
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
    55
functor SimpFun (Simp_data: SIMP_DATA) : SIMP =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
19805
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
    58
local open Simp_data in
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
(*For taking apart reductions into left, right hand sides*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
val lhs_of = #2 o dest_red;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
val rhs_of = #3 o dest_red;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
(*** Indexing and filtering of theorems ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21287
diff changeset
    66
fun eq_brl ((b1 : bool, th1), (b2, th2)) = b1 = b2 andalso Thm.eq_thm_prop (th1, th2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
(*insert a thm in a discrimination net by its lhs*)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
    69
fun lhs_insert_thm th net =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    70
    Net.insert_term eq_brl (lhs_of (Thm.concl_of th), (false,th)) net
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
    handle  Net.INSERT => net;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
(*match subgoal i against possible theorems in the net.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
  Similar to match_from_nat_tac, but the net does not contain numbers;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
  rewrite rules are not ordered.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
fun net_tac net =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
    77
  SUBGOAL(fn (prem,i) =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
    78
          resolve0_tac (Net.unify_term net (Logic.strip_assums_concl prem)) i);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
(*match subgoal i against possible theorems indexed by lhs in the net*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
fun lhs_net_tac net =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
    82
  SUBGOAL(fn (prem,i) =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
    83
          biresolve0_tac (Net.unify_term net
19805
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
    84
                       (lhs_of (Logic.strip_assums_concl prem))) i);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    86
fun nth_subgoal i thm = nth (Thm.prems_of thm) (i - 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
19805
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
    88
fun goal_concl i thm = Logic.strip_assums_concl (nth_subgoal i thm);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
fun lhs_of_eq i thm = lhs_of(goal_concl i thm)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
and rhs_of_eq i thm = rhs_of(goal_concl i thm);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
fun var_lhs(thm,i) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
let fun var(Var _) = true
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
      | var(Abs(_,_,t)) = var t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
      | var(f$_) = var f
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
      | var _ = false;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
in var(lhs_of_eq i thm) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
fun contains_op opns =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36603
diff changeset
   101
    let fun contains(Const(s,_)) = member (op =) opns s |
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
            contains(s$t) = contains s orelse contains t |
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
            contains(Abs(_,_,t)) = contains t |
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
            contains _ = false;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
    in contains end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
fun may_match(match_ops,i) = contains_op match_ops o lhs_of_eq i;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
val (normI_thms,normE_thms) = split_list norm_thms;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
(*Get the norm constants from norm_thms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
val norms =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   113
  let fun norm thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   114
      case lhs_of (Thm.concl_of thm) of
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   115
          Const(n,_)$_ => n
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   116
        | _ => error "No constant in lhs of a norm_thm"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
  in map norm normE_thms end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36603
diff changeset
   120
        Const(s,_)$_ => member (op =) norms s | _ => false;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   122
val refl_tac = resolve0_tac refl_thms;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
fun find_res thms thm =
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   125
    let fun find [] = error "Check Simp_Data"
6969
441393b452c7 handle THM exn;
wenzelm
parents: 5963
diff changeset
   126
          | find(th::thms) = thm RS th handle THM _ => find thms
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
    in find thms end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
val mk_trans = find_res trans_thms;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
fun mk_trans2 thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
let fun mk[] = error"Check transitivity"
6969
441393b452c7 handle THM exn;
wenzelm
parents: 5963
diff changeset
   133
      | mk(t::ts) = (thm RSN (2,t))  handle THM _  => mk ts
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
in mk trans_thms end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   136
(*Applies tactic and returns the first resulting state, FAILS if none!*)
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   137
fun one_result(tac,thm) = case Seq.pull(tac thm) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   138
        SOME(thm',_) => thm'
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   139
      | NONE => raise THM("Simplifier: could not continue", 0, [thm]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   141
fun res1(thm,thms,i) = one_result(resolve0_tac thms i,thm);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
(**** Adding "NORM" tags ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
(*get name of the constant from conclusion of a congruence rule*)
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   147
fun cong_const cong =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   148
    case head_of (lhs_of (Thm.concl_of cong)) of
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   149
        Const(c,_) => c
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   150
      | _ => ""                 (*a placeholder distinct from const names*);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
(*true if the term is an atomic proposition (no ==> signs) *)
19805
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
   153
val atomic = null o Logic.strip_assums_hyp;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
(*ccs contains the names of the constants possessing congruence rules*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
fun add_hidden_vars ccs =
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   157
  let fun add_hvars tm hvars = case tm of
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42364
diff changeset
   158
              Abs(_,_,body) => Misc_Legacy.add_term_vars(body,hvars)
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   159
            | _$_ => let val (f,args) = strip_comb tm
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   160
                     in case f of
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   161
                            Const(c,T) =>
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   162
                                if member (op =) ccs c
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   163
                                then fold_rev add_hvars args hvars
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42364
diff changeset
   164
                                else Misc_Legacy.add_term_vars (tm, hvars)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42364
diff changeset
   165
                          | _ => Misc_Legacy.add_term_vars (tm, hvars)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   166
                     end
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   167
            | _ => hvars;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
  in add_hvars end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
fun add_new_asm_vars new_asms =
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   171
    let fun itf (tm, at) vars =
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42364
diff changeset
   172
                if at then vars else Misc_Legacy.add_term_vars(tm,vars)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   173
        fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   174
                in if length(tml)=length(al)
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   175
                   then fold_rev itf (tml ~~ al) vars
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   176
                   else vars
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   177
                end
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   178
        fun add_vars (tm,vars) = case tm of
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   179
                  Abs (_,_,body) => add_vars(body,vars)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   180
                | r$s => (case head_of tm of
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 16931
diff changeset
   181
                          Const(c,T) => (case AList.lookup (op =) new_asms c of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   182
                                  NONE => add_vars(r,add_vars(s,vars))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   183
                                | SOME(al) => add_list(tm,al,vars))
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   184
                        | _ => add_vars(r,add_vars(s,vars)))
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   185
                | _ => vars
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
    in add_vars end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
fun add_norms(congs,ccs,new_asms) thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
let val thm' = mk_trans2 thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
(* thm': [?z -> l; Prems; r -> ?t] ==> ?z -> ?t *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   192
    val nops = Thm.nprems_of thm'
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
    val lhs = rhs_of_eq 1 thm'
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
    val rhs = lhs_of_eq nops thm'
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   195
    val asms = tl(rev(tl(Thm.prems_of thm')))
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   196
    val hvars = fold_rev (add_hidden_vars ccs) (lhs::rhs::asms) []
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
    val hvars = add_new_asm_vars new_asms (rhs,hvars)
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   198
    fun it_asms asm hvars =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   199
        if atomic asm then add_new_asm_vars new_asms (asm,hvars)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42364
diff changeset
   200
        else Misc_Legacy.add_term_frees(asm,hvars)
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   201
    val hvars = fold_rev it_asms asms hvars
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
    val hvs = map (#1 o dest_Var) hvars
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
    val refl1_tac = refl_tac 1
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   204
    fun norm_step_tac st = st |>
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   205
         (case head_of(rhs_of_eq 1 st) of
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36603
diff changeset
   206
            Var(ixn,_) => if member (op =) hvs ixn then refl1_tac
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   207
                          else resolve0_tac normI_thms 1 ORELSE refl1_tac
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   208
          | Const _ => resolve0_tac normI_thms 1 ORELSE
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   209
                       resolve0_tac congs 1 ORELSE refl1_tac
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59170
diff changeset
   210
          | Free _ => resolve0_tac congs 1 ORELSE refl1_tac
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   211
          | _ => refl1_tac)
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   212
    val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   213
    val SOME(thm'',_) = Seq.pull(add_norm_tac thm')
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
in thm'' end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
fun add_norm_tags congs =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
    let val ccs = map cong_const congs
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   218
        val new_asms = filter (exists not o #2)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   219
                (ccs ~~ (map (map atomic o Thm.prems_of) congs));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
    in add_norms(congs,ccs,new_asms) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
fun normed_rews congs =
19925
3f9341831812 eliminated freeze/varify in favour of Variable.import/export/trade;
wenzelm
parents: 19876
diff changeset
   223
  let val add_norms = add_norm_tags congs in
59170
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   224
    fn thm =>
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   225
      let
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   226
        val ctxt =
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   227
          Thm.theory_of_thm thm
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   228
          |> Proof_Context.init_global
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   229
          |> Variable.declare_thm thm;
de18f8b1a5a2 tuned signature;
wenzelm
parents: 58963
diff changeset
   230
      in Variable.tradeT (K (map (add_norms o mk_trans) o maps mk_rew_rules)) ctxt [thm] end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   231
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   232
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   233
fun NORM norm_lhs_tac = EVERY'[rtac red2 , norm_lhs_tac, refl_tac];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   234
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
val trans_norms = map mk_trans normE_thms;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
(* SIMPSET *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   240
datatype simpset =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   241
        SS of {auto_tac: int -> tactic,
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   242
               congs: thm list,
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   243
               cong_net: thm Net.net,
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   244
               mk_simps: thm -> thm list,
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   245
               simps: (thm * thm list) list,
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   246
               simp_net: thm Net.net}
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
val empty_ss = SS{auto_tac= K no_tac, congs=[], cong_net=Net.empty,
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   249
                  mk_simps=normed_rews[], simps=[], simp_net=Net.empty};
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   251
(** Insertion of congruences and rewrites **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   253
(*insert a thm in a thm net*)
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   254
fun insert_thm_warn th net =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   255
  Net.insert_term Thm.eq_thm_prop (Thm.concl_of th, th) net
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   256
  handle Net.INSERT =>
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   257
    (writeln ("Duplicate rewrite or congruence rule:\n" ^
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   258
        Display.string_of_thm_without_context th); net);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   260
val insert_thms = fold_rev insert_thm_warn;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   261
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33063
diff changeset
   262
fun addrew thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   263
let val thms = mk_simps thm
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
in SS{auto_tac=auto_tac,congs=congs, cong_net=cong_net, mk_simps=mk_simps,
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   265
      simps = (thm,thms)::simps, simp_net = insert_thms thms simp_net}
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   266
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33063
diff changeset
   268
fun ss addrews thms = fold addrew thms ss;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   271
let val congs' = thms @ congs;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   272
in SS{auto_tac=auto_tac, congs= congs',
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   273
      cong_net= insert_thms (map mk_trans thms) cong_net,
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
      mk_simps= normed_rews congs', simps=simps, simp_net=simp_net}
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
(** Deletion of congruences and rewrites **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   278
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
(*delete a thm from a thm net*)
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   280
fun delete_thm_warn th net =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   281
  Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   282
  handle Net.DELETE =>
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   283
    (writeln ("No such rewrite or congruence rule:\n" ^
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   284
        Display.string_of_thm_without_context th); net);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   286
val delete_thms = fold_rev delete_thm_warn;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   287
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21287
diff changeset
   289
let val congs' = fold (remove Thm.eq_thm_prop) thms congs
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
in SS{auto_tac=auto_tac, congs= congs',
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   291
      cong_net= delete_thms (map mk_trans thms) cong_net,
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
      mk_simps= normed_rews congs', simps=simps, simp_net=simp_net}
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33063
diff changeset
   295
fun delrew thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   296
let fun find((p as (th,ths))::ps',ps) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21287
diff changeset
   297
          if Thm.eq_thm_prop(thm,th) then (ths,ps@ps') else find(ps',p::ps)
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   298
      | find([],simps') =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   299
          (writeln ("No such rewrite or congruence rule:\n" ^
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   300
              Display.string_of_thm_without_context thm); ([], simps'))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
    val (thms,simps') = find(simps,[])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
in SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps,
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   303
      simps = simps', simp_net = delete_thms thms simp_net }
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33063
diff changeset
   306
fun ss delrews thms = fold delrew thms ss;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net,...}, auto_tac) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
    SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   311
       simps=simps, simp_net=simp_net};
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   314
(** Inspection of a simpset **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
fun dest_ss(SS{congs,simps,...}) = (congs, map #1 simps);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
fun print_ss(SS{congs,simps,...}) =
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   319
  writeln (cat_lines
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   320
   (["Congruences:"] @ map Display.string_of_thm_without_context congs @
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   321
    ["Rewrite Rules:"] @ map (Display.string_of_thm_without_context o #1) simps));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   322
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   324
(* Rewriting with conditionals *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   325
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
val (case_thms,case_consts) = split_list case_splits;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   327
val case_rews = map mk_trans case_thms;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   328
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   329
fun if_rewritable ifc i thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   330
    let val tm = goal_concl i thm
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   331
        fun nobound(Abs(_,_,tm),j,k) = nobound(tm,j,k+1)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   332
          | nobound(s$t,j,k) = nobound(s,j,k) andalso nobound(t,j,k)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   333
          | nobound(Bound n,j,k) = n < k orelse k+j <= n
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   334
          | nobound(_) = true;
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   335
        fun check_args(al,j) = forall (fn t => nobound(t,j,0)) al
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   336
        fun find_if(Abs(_,_,tm),j) = find_if(tm,j+1)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   337
          | find_if(tm as s$t,j) = let val (f,al) = strip_comb tm in
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   338
                case f of Const(c,_) => if c=ifc then check_args(al,j)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   339
                        else find_if(s,j) orelse find_if(t,j)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   340
                | _ => find_if(s,j) orelse find_if(t,j) end
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   341
          | find_if(_) = false;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   342
    in find_if(tm,0) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   343
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   344
fun IF1_TAC cong_tac i =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   345
    let fun seq_try (ifth::ifths,ifc::ifcs) thm =
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   346
                (COND (if_rewritable ifc i) (DETERM(rtac ifth i))
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   347
                        (seq_try(ifths,ifcs))) thm
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   348
              | seq_try([],_) thm = no_tac thm
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   349
        and try_rew thm = (seq_try(case_rews,case_consts) ORELSE one_subt) thm
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   350
        and one_subt thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   351
                let val test = has_fewer_prems (Thm.nprems_of thm + 1)
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   352
                    fun loop thm =
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   353
                        COND test no_tac
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   354
                          ((try_rew THEN DEPTH_FIRST test (refl_tac i))
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   355
                           ORELSE (refl_tac i THEN loop)) thm
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   356
                in (cong_tac THEN loop) thm end
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   357
    in COND (may_match(case_consts,i)) try_rew no_tac end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   358
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   359
fun CASE_TAC (SS{cong_net,...}) i =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   360
let val cong_tac = net_tac cong_net i
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   361
in NORM (IF1_TAC cong_tac) i end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   362
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   363
(* Rewriting Automaton *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
datatype cntrl = STOP | MK_EQ | ASMS of int | SIMP_LHS | REW | REFL | TRUE
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   366
               | PROVE | POP_CS | POP_ARTR | IF;
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22360
diff changeset
   367
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   368
fun simp_refl([],_,ss) = ss
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   369
  | simp_refl(a'::ns,a,ss) = if a'=a then simp_refl(ns,a,SIMP_LHS::REFL::ss)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   370
        else simp_refl(ns,a,ASMS(a)::SIMP_LHS::REFL::POP_ARTR::ss);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
(** Tracing **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   373
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32449
diff changeset
   374
val tracing = Unsynchronized.ref false;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   375
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
(*Replace parameters by Free variables in P*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   377
fun variants_abs ([],P) = P
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   378
  | variants_abs ((a,T)::aTs, P) =
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 36945
diff changeset
   379
      variants_abs (aTs, #2 (Syntax_Trans.variant_abs(a,T,P)));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   380
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   381
(*Select subgoal i from proof state; substitute parameters, for printing*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   382
fun prepare_goal i st =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   383
    let val subgi = nth_subgoal i st
19805
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
   384
        val params = rev (Logic.strip_params subgi)
854404b8f738 do not open Logic;
wenzelm
parents: 17325
diff changeset
   385
    in variants_abs (params, Logic.strip_assums_concl subgi) end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   386
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   387
(*print lhs of conclusion of subgoal i*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   388
fun pr_goal_lhs i st =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   389
    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   390
             (lhs_of (prepare_goal i st)));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   391
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   392
(*print conclusion of subgoal i*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   393
fun pr_goal_concl i st =
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   394
    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st) (prepare_goal i st))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   395
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   396
(*print subgoals i to j (inclusive)*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   397
fun pr_goals (i,j) st =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   398
    if i>j then ()
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   399
    else (pr_goal_concl i st;  pr_goals (i+1,j) st);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   400
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   401
(*Print rewrite for tracing; i=subgoal#, n=number of new subgoals,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   402
  thm=old state, thm'=new state *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   403
fun pr_rew (i,n,thm,thm',not_asms) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   404
    if !tracing
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   405
    then (if not_asms then () else writeln"Assumption used in";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   406
          pr_goal_lhs i thm; writeln"->"; pr_goal_lhs (i+n) thm';
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   407
          if n>0 then (writeln"Conditions:"; pr_goals (i, i+n-1) thm')
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
          else ();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   409
          writeln"" )
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   410
    else ();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   411
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   412
(* Skip the first n hyps of a goal, and return the rest in generalized form *)
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 44121
diff changeset
   413
fun strip_varify(Const(@{const_name Pure.imp}, _) $ H $ B, n, vs) =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   414
        if n=0 then subst_bounds(vs,H)::strip_varify(B,0,vs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   415
        else strip_varify(B,n-1,vs)
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 44121
diff changeset
   416
  | strip_varify(Const(@{const_name Pure.all},_)$Abs(_,T,t), n, vs) =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   417
        strip_varify(t,n,Var(("?",length vs),T)::vs)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   418
  | strip_varify  _  = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   419
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   420
fun execute(ss,if_fl,auto_tac,cong_tac,net,i,thm) = let
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   422
fun simp_lhs(thm,ss,anet,ats,cs) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
    if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   424
    if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs)
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   425
    else case Seq.pull(cong_tac i thm) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   426
            SOME(thm',_) =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   427
                    let val ps = Thm.prems_of thm
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   428
                        and ps' = Thm.prems_of thm';
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   429
                        val n = length(ps')-length(ps);
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42284
diff changeset
   430
                        val a = length(Logic.strip_assums_hyp(nth ps (i - 1)))
33955
fff6f11b1f09 curried take/drop
haftmann
parents: 33339
diff changeset
   431
                        val l = map (length o Logic.strip_assums_hyp) (take n (drop (i-1) ps'));
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   432
                    in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   433
          | NONE => (REW::ss,thm,anet,ats,cs);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   434
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   435
(*NB: the "Adding rewrites:" trace will look strange because assumptions
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   436
      are represented by rules, generalized over their parameters*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   437
fun add_asms(ss,thm,a,anet,ats,cs) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   438
    let val As = strip_varify(nth_subgoal i thm, a, []);
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   439
        val thms = map (Thm.trivial o Thm.global_cterm_of(Thm.theory_of_thm thm)) As;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   440
        val new_rws = maps mk_rew_rules thms;
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   441
        val rwrls = map mk_trans (maps mk_rew_rules thms);
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   442
        val anet' = fold_rev lhs_insert_thm rwrls anet;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   443
    in  if !tracing andalso not(null new_rws)
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   444
        then writeln (cat_lines
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30190
diff changeset
   445
          ("Adding rewrites:" :: map Display.string_of_thm_without_context new_rws))
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   446
        else ();
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   447
        (ss,thm,anet',anet::ats,cs)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   448
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   449
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   450
fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   451
      SOME(thm',seq') =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   452
            let val n = (Thm.nprems_of thm') - (Thm.nprems_of thm)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   453
            in pr_rew(i,n,thm,thm',more);
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   454
               if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   455
               else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss),
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   456
                     thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   457
            end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   458
    | NONE => if more
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 56245
diff changeset
   459
            then rew((lhs_net_tac anet i THEN atac i) thm,
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   460
                     thm,ss,anet,ats,cs,false)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   461
            else (ss,thm,anet,ats,cs);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   462
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   463
fun try_true(thm,ss,anet,ats,cs) =
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   464
    case Seq.pull(auto_tac i thm) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   465
      SOME(thm',_) => (ss,thm',anet,ats,cs)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   466
    | NONE => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   467
              in if !tracing
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   468
                 then (writeln"*** Failed to prove precondition. Normal form:";
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   469
                       pr_goal_concl i thm;  writeln"")
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   470
                 else ();
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   471
                 rew(seq,thm0,ss0,anet0,ats0,cs0,more)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   472
              end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   473
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
fun if_exp(thm,ss,anet,ats,cs) =
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   475
        case Seq.pull (IF1_TAC (cong_tac i) i thm) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   476
                SOME(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   477
              | NONE => (ss,thm,anet,ats,cs);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   478
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   479
fun step(s::ss, thm, anet, ats, cs) = case s of
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   480
          MK_EQ => (ss, res1(thm,[red2],i), anet, ats, cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   481
        | ASMS(a) => add_asms(ss,thm,a,anet,ats,cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   482
        | SIMP_LHS => simp_lhs(thm,ss,anet,ats,cs)
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   483
        | REW => rew(net_tac net i thm,thm,ss,anet,ats,cs,true)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   484
        | REFL => (ss, res1(thm,refl_thms,i), anet, ats, cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   485
        | TRUE => try_true(res1(thm,refl_thms,i),ss,anet,ats,cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   486
        | PROVE => (if if_fl then MK_EQ::SIMP_LHS::IF::TRUE::ss
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   487
                    else MK_EQ::SIMP_LHS::TRUE::ss, thm, anet, ats, cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   488
        | POP_ARTR => (ss,thm,hd ats,tl ats,cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   489
        | POP_CS => (ss,thm,anet,ats,tl cs)
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   490
        | IF => if_exp(thm,ss,anet,ats,cs);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   491
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   492
fun exec(state as (s::ss, thm, _, _, _)) =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   493
        if s=STOP then thm else exec(step(state));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   494
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   495
in exec(ss, thm, Net.empty, [], []) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   496
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   497
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   498
fun EXEC_TAC(ss,fl) (SS{auto_tac,cong_net,simp_net,...}) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   499
let val cong_tac = net_tac cong_net
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 32091
diff changeset
   500
in fn i =>
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   501
    (fn thm =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   502
     if i <= 0 orelse Thm.nprems_of thm < i then Seq.empty
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 3537
diff changeset
   503
     else Seq.single(execute(ss,fl,auto_tac,cong_tac,simp_net,i,thm)))
1512
ce37c64244c0 Elimination of fully-functorial style.
paulson
parents: 1459
diff changeset
   504
    THEN TRY(auto_tac i)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   505
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   506
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   507
val SIMP_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,REFL,STOP],false);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   508
val SIMP_CASE_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,IF,REFL,STOP],false);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   509
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   510
val ASM_SIMP_TAC = EXEC_TAC([ASMS(0),MK_EQ,SIMP_LHS,REFL,STOP],false);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   511
val ASM_SIMP_CASE_TAC = EXEC_TAC([ASMS(0),MK_EQ,SIMP_LHS,IF,REFL,STOP],false);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   512
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   513
val SIMP_CASE2_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,IF,REFL,STOP],true);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   514
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   515
fun REWRITE (ss,fl) (SS{auto_tac,cong_net,simp_net,...}) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   516
let val cong_tac = net_tac cong_net
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   517
in fn thm => let val state = thm RSN (2,red1)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   518
             in execute(ss,fl,auto_tac,cong_tac,simp_net,1,state) end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   519
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   520
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   521
val SIMP_THM = REWRITE ([ASMS(0),SIMP_LHS,IF,REFL,STOP],false);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   522
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   523
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   524
(* Compute Congruence rules for individual constants using the substition
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   525
   rules *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   526
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33955
diff changeset
   527
val subst_thms = map Drule.export_without_context subst_thms;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   528
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   529
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   530
fun exp_app(0,t) = t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   531
  | exp_app(i,t) = exp_app(i-1,t $ Bound (i-1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   532
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   533
fun exp_abs(Type("fun",[T1,T2]),t,i) =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   534
        Abs("x"^string_of_int i,T1,exp_abs(T2,t,i+1))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   535
  | exp_abs(T,t,i) = exp_app(i,t);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   536
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   537
fun eta_Var(ixn,T) = exp_abs(T,Var(ixn,T),0);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   538
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   539
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   540
fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   541
let fun xn_list(x,n) =
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33040
diff changeset
   542
        let val ixs = map_range (fn i => (x^(radixstring(26,"a",i)),0)) (n - 1);
33955
fff6f11b1f09 curried take/drop
haftmann
parents: 33339
diff changeset
   543
        in ListPair.map eta_Var (ixs, take (n+1) Ts) end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   544
    val lhs = list_comb(f,xn_list("X",k-1))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   545
    val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   546
in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   547
16931
e41d8e319dfd Sign.typ_instance;
wenzelm
parents: 16876
diff changeset
   548
fun find_subst sg T =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   549
let fun find (thm::thms) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   550
        let val (Const(_,cT), va, vb) = dest_red(hd(Thm.prems_of thm));
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   551
            val [P] = subtract (op =) [va, vb] (Misc_Legacy.add_term_vars (Thm.concl_of thm, []));
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   552
            val eqT::_ = binder_types cT
16931
e41d8e319dfd Sign.typ_instance;
wenzelm
parents: 16876
diff changeset
   553
        in if Sign.typ_instance sg (T,eqT) then SOME(thm,va,vb,P)
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   554
           else find thms
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   555
        end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   556
      | find [] = NONE
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   557
in find subst_thms end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   558
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   559
fun mk_cong sg (f,aTs,rT) (refl,eq) =
16931
e41d8e319dfd Sign.typ_instance;
wenzelm
parents: 16876
diff changeset
   560
let val k = length aTs;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   561
    fun ri((subst,va as Var(_,Ta),vb as Var(_,Tb),P),i,si,T,yik) =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   562
        let val ca = Thm.global_cterm_of sg va
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   563
            and cx = Thm.global_cterm_of sg (eta_Var(("X"^si,0),T))
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   564
            val cb = Thm.global_cterm_of sg vb
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   565
            and cy = Thm.global_cterm_of sg (eta_Var(("Y"^si,0),T))
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   566
            val cP = Thm.global_cterm_of sg P
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   567
            and cp = Thm.global_cterm_of sg (Pinst(f,rT,eq,k,i,T,yik,aTs))
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   568
        in cterm_instantiate [(ca,cx),(cb,cy),(cP,cp)] subst end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   569
    fun mk(c,T::Ts,i,yik) =
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   570
        let val si = radixstring(26,"a",i)
16931
e41d8e319dfd Sign.typ_instance;
wenzelm
parents: 16876
diff changeset
   571
        in case find_subst sg T of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   572
             NONE => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   573
           | SOME s => let val c' = c RSN (2,ri(s,i,si,T,yik))
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   574
                       in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   575
        end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   576
      | mk(c,[],_,_) = c;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   577
in mk(refl,rev aTs,k-1,[]) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   578
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   579
fun mk_cong_type sg (f,T) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   580
let val (aTs,rT) = strip_type T;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   581
    fun find_refl(r::rs) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   582
        let val (Const(eq,eqT),_,_) = dest_red(Thm.concl_of r)
16931
e41d8e319dfd Sign.typ_instance;
wenzelm
parents: 16876
diff changeset
   583
        in if Sign.typ_instance sg (rT, hd(binder_types eqT))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   584
           then SOME(r,(eq,body_type eqT)) else find_refl rs
1459
d12da312eff4 expanded tabs
clasohm
parents: 611
diff changeset
   585
        end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   586
      | find_refl([]) = NONE;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   587
in case find_refl refl_thms of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   588
     NONE => []  |  SOME(refl) => [mk_cong sg (f,aTs,rT) refl]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   589
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   590
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   591
fun mk_cong_thy thy f =
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22360
diff changeset
   592
let val T = case Sign.const_type thy f of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14772
diff changeset
   593
                NONE => error(f^" not declared") | SOME(T) => T;
16876
f57b38cced32 Logic.incr_tvar;
wenzelm
parents: 16800
diff changeset
   594
    val T' = Logic.incr_tvar 9 T;
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22360
diff changeset
   595
in mk_cong_type thy (Const(f,T'),T') end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   596
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   597
fun mk_congs thy = maps (mk_cong_thy thy);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   598
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   599
fun mk_typed_congs thy =
22675
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   600
let
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   601
  fun readfT(f,s) =
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   602
    let
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 22675
diff changeset
   603
      val T = Logic.incr_tvar 9 (Syntax.read_typ_global thy s);
22675
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   604
      val t = case Sign.const_type thy f of
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   605
                  SOME(_) => Const(f,T) | NONE => Free(f,T)
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   606
    in (t,T) end
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   607
in maps (mk_cong_type thy o readfT) end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   608
22675
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   609
end;
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22596
diff changeset
   610
end;