src/HOL/Tools/Function/scnp_reconstruct.ML
author wenzelm
Thu, 11 Aug 2016 18:26:16 +0200
changeset 63668 5efaa884ac6c
parent 63170 eae6549dbea2
child 69593 3dda49e08b9d
permissions -rw-r--r--
tuned error;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31242
diff changeset
     1
(*  Title:       HOL/Tools/Function/scnp_reconstruct.ML
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     2
    Author:      Armin Heller, TU Muenchen
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     3
    Author:      Alexander Krauss, TU Muenchen
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     4
41114
f9ae7c2abf7e tuned headers
krauss
parents: 40317
diff changeset
     5
Proof reconstruction for SCNP termination.
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     6
*)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     7
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     8
signature SCNP_RECONSTRUCT =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     9
sig
29877
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
    10
  val sizechange_tac : Proof.context -> tactic -> tactic
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
    11
36521
73ed9f18fdd3 default termination prover as plain tactic
krauss
parents: 35624
diff changeset
    12
  val decomp_scnp_tac : ScnpSolve.label list -> Proof.context -> tactic
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    13
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    14
  datatype multiset_setup =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    15
    Multiset of
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    16
    {
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    17
     msetT : typ -> typ,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    18
     mk_mset : typ -> term list -> term,
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
    19
     mset_regroup_conv : Proof.context -> int list -> conv,
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    20
     mset_member_tac : Proof.context -> int -> int -> tactic,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    21
     mset_nonempty_tac : Proof.context -> int -> tactic,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    22
     mset_pwleq_tac : Proof.context -> int -> tactic,
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    23
     set_of_simps : thm list,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    24
     smsI' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    25
     wmsI2'' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    26
     wmsI1 : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    27
     reduction_pair : thm
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    28
    }
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    29
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    30
  val multiset_setup : multiset_setup -> theory -> theory
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    31
end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    32
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    33
structure ScnpReconstruct : SCNP_RECONSTRUCT =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    34
struct
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    35
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33063
diff changeset
    36
val PROFILE = Function_Common.PROFILE
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    37
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    38
open ScnpSolve
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    39
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    40
val natT = HOLogic.natT
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    41
val nat_pairT = HOLogic.mk_prodT (natT, natT)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    42
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    43
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    44
(* Theory dependencies *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    45
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    46
datatype multiset_setup =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    47
  Multiset of
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    48
  {
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    49
   msetT : typ -> typ,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    50
   mk_mset : typ -> term list -> term,
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
    51
   mset_regroup_conv : Proof.context -> int list -> conv,
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    52
   mset_member_tac : Proof.context -> int -> int -> tactic,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    53
   mset_nonempty_tac : Proof.context -> int -> tactic,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
    54
   mset_pwleq_tac : Proof.context -> int -> tactic,
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    55
   set_of_simps : thm list,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    56
   smsI' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    57
   wmsI2'' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    58
   wmsI1 : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    59
   reduction_pair : thm
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    60
  }
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    61
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33468
diff changeset
    62
structure Multiset_Setup = Theory_Data
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    63
(
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    64
  type T = multiset_setup option
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    65
  val empty = NONE
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    66
  val extend = I;
41493
f05976d69141 added merge_options convenience;
wenzelm
parents: 41114
diff changeset
    67
  val merge = merge_options
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    68
)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    69
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33063
diff changeset
    70
val multiset_setup = Multiset_Setup.put o SOME
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    71
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33583
diff changeset
    72
fun undef _ = error "undef"
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    73
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
    74
fun get_multiset_setup ctxt = Multiset_Setup.get (Proof_Context.theory_of ctxt)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    75
  |> the_default (Multiset
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    76
    { msetT = undef, mk_mset=undef,
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    77
      mset_regroup_conv=undef, mset_member_tac = undef,
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    78
      mset_nonempty_tac = undef, mset_pwleq_tac = undef,
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    79
      set_of_simps = [],reduction_pair = refl,
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    80
      smsI'=refl, wmsI2''=refl, wmsI1=refl })
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    81
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    82
fun order_rpair _ MAX = @{thm max_rpair_set}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    83
  | order_rpair msrp MS  = msrp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    84
  | order_rpair _ MIN = @{thm min_rpair_set}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    85
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    86
fun ord_intros_max true = (@{thm smax_emptyI}, @{thm smax_insertI})
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    87
  | ord_intros_max false = (@{thm wmax_emptyI}, @{thm wmax_insertI})
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    88
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    89
fun ord_intros_min true = (@{thm smin_emptyI}, @{thm smin_insertI})
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
    90
  | ord_intros_min false = (@{thm wmin_emptyI}, @{thm wmin_insertI})
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    91
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    92
fun gen_probl D cs =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    93
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    94
    val n = Termination.get_num_points D
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    95
    val arity = length o Termination.get_measures D
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    96
    fun measure p i = nth (Termination.get_measures D p) i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    97
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    98
    fun mk_graph c =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    99
      let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   100
        val (_, p, _, q, _, _) = Termination.dest_call D c
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   101
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   102
        fun add_edge i j =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   103
          case Termination.get_descent D c (measure p i) (measure q j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   104
           of SOME (Termination.Less _) => cons (i, GTR, j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   105
            | SOME (Termination.LessEq _) => cons (i, GEQ, j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   106
            | _ => I
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   107
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   108
        val edges =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   109
          fold_product add_edge (0 upto arity p - 1) (0 upto arity q - 1) []
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   110
      in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   111
        G (p, q, edges)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   112
      end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   113
  in
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33040
diff changeset
   114
    GP (map_range arity n, map mk_graph cs)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   115
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   116
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   117
(* General reduction pair application *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   118
fun rem_inv_img ctxt =
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   119
  resolve_tac ctxt @{thms subsetI} 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   120
  THEN eresolve_tac ctxt @{thms CollectE} 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   121
  THEN REPEAT (eresolve_tac ctxt @{thms exE} 1)
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 60801
diff changeset
   122
  THEN Local_Defs.unfold0_tac ctxt @{thms inv_image_def}
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   123
  THEN resolve_tac ctxt @{thms CollectI} 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   124
  THEN eresolve_tac ctxt @{thms conjE} 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   125
  THEN eresolve_tac ctxt @{thms ssubst} 1
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 60801
diff changeset
   126
  THEN Local_Defs.unfold0_tac ctxt @{thms split_conv triv_forall_equality sum.case}
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   127
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   128
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   129
(* Sets *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   130
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   131
val setT = HOLogic.mk_setT
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   132
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   133
fun set_member_tac ctxt m i =
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   134
  if m = 0 then resolve_tac ctxt @{thms insertI1} i
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   135
  else resolve_tac ctxt @{thms insertI2} i THEN set_member_tac ctxt (m - 1) i
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   136
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   137
fun set_nonempty_tac ctxt = resolve_tac ctxt @{thms insert_not_empty}
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   138
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   139
fun set_finite_tac ctxt i =
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   140
  resolve_tac ctxt @{thms finite.emptyI} i
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   141
  ORELSE (resolve_tac ctxt @{thms finite.insertI} i THEN (fn st => set_finite_tac ctxt i st))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   142
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   143
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   144
(* Reconstruction *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   145
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33583
diff changeset
   146
fun reconstruct_tac ctxt D cs (GP (_, gs)) certificate =
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   147
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   148
    val Multiset
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   149
          { msetT, mk_mset,
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33583
diff changeset
   150
            mset_regroup_conv, mset_pwleq_tac, set_of_simps,
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   151
            smsI', wmsI2'', wmsI1, reduction_pair=ms_rp, ...}
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   152
        = get_multiset_setup ctxt
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   153
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   154
    fun measure_fn p = nth (Termination.get_measures D p)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   155
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   156
    fun get_desc_thm cidx m1 m2 bStrict =
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   157
      (case Termination.get_descent D (nth cs cidx) m1 m2 of
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   158
        SOME (Termination.Less thm) =>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   159
          if bStrict then thm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   160
          else (thm COMP (Thm.lift_rule (Thm.cprop_of thm) @{thm less_imp_le}))
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   161
      | SOME (Termination.LessEq (thm, _))  =>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   162
          if not bStrict then thm
40317
1eac228c52b3 replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
wenzelm
parents: 39925
diff changeset
   163
          else raise Fail "get_desc_thm"
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   164
      | _ => raise Fail "get_desc_thm")
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   165
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   166
    val (label, lev, sl, covering) = certificate
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   167
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   168
    fun prove_lev strict g =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   169
      let
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33583
diff changeset
   170
        val G (p, q, _) = nth gs g
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   171
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   172
        fun less_proof strict (j, b) (i, a) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   173
          let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   174
            val tag_flag = b < a orelse (not strict andalso b <= a)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   175
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   176
            val stored_thm =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   177
              get_desc_thm g (measure_fn p i) (measure_fn q j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   178
                             (not tag_flag)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   179
              |> Conv.fconv_rule (Thm.beta_conversion true)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   180
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   181
            val rule =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   182
              if strict
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   183
              then if b < a then @{thm pair_lessI2} else @{thm pair_lessI1}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   184
              else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   185
          in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   186
            resolve_tac ctxt [rule] 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
33569
1ebb8b7b9f6a make "sizechange_tac" slightly less verbose
blanchet
parents: 33099
diff changeset
   187
            THEN (if tag_flag then Arith_Data.arith_tac ctxt 1 else all_tac)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   188
          end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   189
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   190
        fun steps_tac MAX strict lq lp =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   191
              let
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   192
                val (empty, step) = ord_intros_max strict
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   193
              in
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   194
                if length lq = 0
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   195
                then resolve_tac ctxt [empty] 1 THEN set_finite_tac ctxt 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   196
                     THEN (if strict then set_nonempty_tac ctxt 1 else all_tac)
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   197
                else
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   198
                  let
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   199
                    val (j, b) :: rest = lq
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   200
                    val (i, a) = the (covering g strict j)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   201
                    fun choose xs = set_member_tac ctxt (find_index (curry op = (i, a)) xs) 1
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   202
                    val solve_tac = choose lp THEN less_proof strict (j, b) (i, a)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   203
                  in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   204
                    resolve_tac ctxt [step] 1 THEN solve_tac THEN steps_tac MAX strict rest lp
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   205
                  end
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   206
              end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   207
          | steps_tac MIN strict lq lp =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   208
              let
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   209
                val (empty, step) = ord_intros_min strict
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   210
              in
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   211
                if length lp = 0
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   212
                then resolve_tac ctxt [empty] 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   213
                     THEN (if strict then set_nonempty_tac ctxt 1 else all_tac)
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   214
                else
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   215
                  let
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   216
                    val (i, a) :: rest = lp
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   217
                    val (j, b) = the (covering g strict i)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   218
                    fun choose xs = set_member_tac ctxt (find_index (curry op = (j, b)) xs) 1
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   219
                    val solve_tac = choose lq THEN less_proof strict (j, b) (i, a)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   220
                  in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   221
                    resolve_tac ctxt [step] 1 THEN solve_tac THEN steps_tac MIN strict lq rest
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   222
                  end
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   223
              end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   224
          | steps_tac MS strict lq lp =
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   225
              let
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   226
                fun get_str_cover (j, b) =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   227
                  if is_some (covering g true j) then SOME (j, b) else NONE
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   228
                fun get_wk_cover (j, b) = the (covering g false j)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   229
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   230
                val qs = subtract (op =) (map_filter get_str_cover lq) lq
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   231
                val ps = map get_wk_cover qs
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   232
59584
wenzelm
parents: 59582
diff changeset
   233
                fun indices xs ys = map (fn y => find_index (curry op = y) xs) ys
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   234
                val iqs = indices lq qs
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   235
                val ips = indices lp ps
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   236
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   237
                local open Conv in
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   238
                fun t_conv a C =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   239
                  params_conv ~1 (K ((concl_conv ~1 o arg_conv o arg1_conv o a) C)) ctxt
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   240
                val goal_rewrite =
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   241
                    t_conv arg1_conv (mset_regroup_conv ctxt iqs)
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   242
                    then_conv t_conv arg_conv (mset_regroup_conv ctxt ips)
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   243
                end
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   244
              in
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   245
                CONVERSION goal_rewrite 1
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   246
                THEN (if strict then resolve_tac ctxt [smsI'] 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   247
                      else if qs = lq then resolve_tac ctxt [wmsI2''] 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   248
                      else resolve_tac ctxt [wmsI1] 1)
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   249
                THEN mset_pwleq_tac ctxt 1
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   250
                THEN EVERY (map2 (less_proof false) qs ps)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   251
                THEN (if strict orelse qs <> lq
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 60801
diff changeset
   252
                      then Local_Defs.unfold0_tac ctxt set_of_simps
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   253
                           THEN steps_tac MAX true
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   254
                           (subtract (op =) qs lq) (subtract (op =) ps lp)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   255
                      else all_tac)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   256
              end
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   257
      in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   258
        rem_inv_img ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   259
        THEN steps_tac label strict (nth lev q) (nth lev p)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   260
      end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   261
30450
7655e6533209 HOLogic.mk_set, HOLogic.dest_set
haftmann
parents: 30304
diff changeset
   262
    val (mk_set, setT) = if label = MS then (mk_mset, msetT) else (HOLogic.mk_set, setT)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   263
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   264
    fun tag_pair p (i, tag) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   265
      HOLogic.pair_const natT natT $
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   266
        (measure_fn p i $ Bound 0) $ HOLogic.mk_number natT tag
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   267
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   268
    fun pt_lev (p, lm) =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   269
      Abs ("x", Termination.get_types D p, mk_set nat_pairT (map (tag_pair p) lm))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   270
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   271
    val level_mapping =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   272
      map_index pt_lev lev
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   273
        |> Termination.mk_sumcases D (setT nat_pairT)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   274
        |> Thm.cterm_of ctxt
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   275
    in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   276
      PROFILE "Proof Reconstruction"
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   277
        (CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv ctxt sl))) 1
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   278
         THEN (resolve_tac ctxt @{thms reduction_pair_lemma} 1)
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   279
         THEN (resolve_tac ctxt @{thms rp_inv_image_rp} 1)
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   280
         THEN (resolve_tac ctxt [order_rpair ms_rp label] 1)
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   281
         THEN PRIMITIVE (Thm.instantiate' [] [SOME level_mapping])
54998
8601434fa334 tuned signature;
wenzelm
parents: 51717
diff changeset
   282
         THEN unfold_tac ctxt @{thms rp_inv_image_def}
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 60801
diff changeset
   283
         THEN Local_Defs.unfold0_tac ctxt @{thms split_conv fst_conv snd_conv}
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59159
diff changeset
   284
         THEN REPEAT (SOMEGOAL (resolve_tac ctxt [@{thm Un_least}, @{thm empty_subsetI}]))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   285
         THEN EVERY (map (prove_lev true) sl)
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   286
         THEN EVERY (map (prove_lev false) (subtract (op =) sl (0 upto length cs - 1))))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   287
    end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   288
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   289
39924
f4d3e70ed3a8 discontinued continuations to simplify control flow; dropped optimization in scnp
krauss
parents: 39923
diff changeset
   290
fun single_scnp_tac use_tags orders ctxt D = Termination.CALLS (fn (cs, i) =>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   291
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41493
diff changeset
   292
    val ms_configured = is_some (Multiset_Setup.get (Proof_Context.theory_of ctxt))
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   293
    val orders' =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   294
      if ms_configured then orders
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   295
      else filter_out (curry op = MS) orders
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   296
    val gp = gen_probl D cs
29877
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
   297
    val certificate = generate_certificate use_tags orders' gp
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
   298
  in
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   299
    (case certificate of
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   300
      NONE => no_tac
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   301
    | SOME cert =>
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   302
        SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   303
        THEN TRY (resolve_tac ctxt @{thms wf_empty} i))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   304
  end)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   305
39924
f4d3e70ed3a8 discontinued continuations to simplify control flow; dropped optimization in scnp
krauss
parents: 39923
diff changeset
   306
fun gen_decomp_scnp_tac orders autom_tac ctxt =
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   307
  Termination.TERMINATION ctxt autom_tac (fn D =>
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   308
    let
59618
e6939796717e clarified context;
wenzelm
parents: 59584
diff changeset
   309
      val decompose = Termination.decompose_tac ctxt D
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   310
      val scnp_full = single_scnp_tac true orders ctxt D
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   311
    in
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   312
      REPEAT_ALL_NEW (scnp_full ORELSE' decompose)
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   313
    end)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   314
39924
f4d3e70ed3a8 discontinued continuations to simplify control flow; dropped optimization in scnp
krauss
parents: 39923
diff changeset
   315
fun gen_sizechange_tac orders autom_tac ctxt =
59159
9312710451f5 just one data slot per program unit;
wenzelm
parents: 58819
diff changeset
   316
  TRY (Function_Common.termination_rule_tac ctxt 1)
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30510
diff changeset
   317
  THEN TRY (Termination.wf_union_tac ctxt)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59625
diff changeset
   318
  THEN (resolve_tac ctxt @{thms wf_empty} 1 ORELSE gen_decomp_scnp_tac orders autom_tac ctxt 1)
29877
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
   319
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
   320
fun sizechange_tac ctxt autom_tac =
39924
f4d3e70ed3a8 discontinued continuations to simplify control flow; dropped optimization in scnp
krauss
parents: 39923
diff changeset
   321
  gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt
29877
867a0ed7a9b2 Export tactic interface for sizechange method
krauss
parents: 29183
diff changeset
   322
36521
73ed9f18fdd3 default termination prover as plain tactic
krauss
parents: 35624
diff changeset
   323
fun decomp_scnp_tac orders ctxt =
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   324
  let
57959
1bfed12a7646 updated to named_theorems;
wenzelm
parents: 55642
diff changeset
   325
    val extra_simps = Named_Theorems.get ctxt @{named_theorems termination_simp}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 42795
diff changeset
   326
    val autom_tac = auto_tac (ctxt addsimps extra_simps)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   327
  in
39924
f4d3e70ed3a8 discontinued continuations to simplify control flow; dropped optimization in scnp
krauss
parents: 39923
diff changeset
   328
     gen_sizechange_tac orders autom_tac ctxt
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   329
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   330
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   331
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   332
(* Method setup *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   333
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   334
val orders =
31242
ed40b987a474 modernized method setup;
wenzelm
parents: 30686
diff changeset
   335
  Scan.repeat1
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   336
    ((Args.$$$ "max" >> K MAX) ||
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   337
     (Args.$$$ "min" >> K MIN) ||
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   338
     (Args.$$$ "ms" >> K MS))
31242
ed40b987a474 modernized method setup;
wenzelm
parents: 30686
diff changeset
   339
  || Scan.succeed [MAX, MS, MIN]
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   340
58819
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   341
val _ =
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   342
  Theory.setup
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   343
    (Method.setup @{binding size_change}
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   344
      (Scan.lift orders --| Method.sections clasimp_modifiers >>
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   345
        (fn orders => SIMPLE_METHOD o decomp_scnp_tac orders))
aa43c6f05bca modernized setup;
wenzelm
parents: 57959
diff changeset
   346
      "termination prover with graph decomposition and the NP subset of size change termination")
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   347
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   348
end