src/HOL/Tools/function_package/scnp_reconstruct.ML
author krauss
Fri, 02 Jan 2009 22:06:56 +0100
changeset 29329 e02b3a32f34f
parent 29183 f1648e009dc1
child 29877 867a0ed7a9b2
permissions -rw-r--r--
removed references to OldTerm.*
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     1
(*  Title:       HOL/Tools/function_package/scnp_reconstruct.ML
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
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
     5
Proof reconstruction for SCNP
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
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    10
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    11
  val decomp_scnp : ScnpSolve.label list -> Proof.context -> method
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    12
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    13
  val setup : theory -> theory
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    14
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    15
  datatype multiset_setup =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    16
    Multiset of
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    17
    {
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    18
     msetT : typ -> typ,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    19
     mk_mset : typ -> term list -> term,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    20
     mset_regroup_conv : int list -> conv,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    21
     mset_member_tac : int -> int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    22
     mset_nonempty_tac : int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    23
     mset_pwleq_tac : int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    24
     set_of_simps : thm list,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    25
     smsI' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    26
     wmsI2'' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    27
     wmsI1 : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    28
     reduction_pair : thm
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
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    31
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    32
  val multiset_setup : multiset_setup -> theory -> theory
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    33
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    34
end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    35
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    36
structure ScnpReconstruct : SCNP_RECONSTRUCT =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    37
struct
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    38
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    39
val PROFILE = FundefCommon.PROFILE
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    40
fun TRACE x = if ! FundefCommon.profile then Output.tracing x else ()
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    41
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    42
open ScnpSolve
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    43
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    44
val natT = HOLogic.natT
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    45
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
    46
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    47
(* Theory dependencies *)
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
datatype multiset_setup =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    50
  Multiset of
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    51
  {
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    52
   msetT : typ -> typ,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    53
   mk_mset : typ -> term list -> term,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    54
   mset_regroup_conv : int list -> conv,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    55
   mset_member_tac : int -> int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    56
   mset_nonempty_tac : int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    57
   mset_pwleq_tac : int -> tactic,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    58
   set_of_simps : thm list,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    59
   smsI' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    60
   wmsI2'' : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    61
   wmsI1 : thm,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    62
   reduction_pair : thm
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
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    65
structure MultisetSetup = TheoryDataFun
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    66
(
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    67
  type T = multiset_setup option
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    68
  val empty = NONE
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    69
  val copy = I;
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    70
  val extend = I;
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    71
  fun merge _ (v1, v2) = if is_some v2 then v2 else v1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    72
)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    73
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    74
val multiset_setup = MultisetSetup.put o SOME
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    75
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    76
fun undef x = error "undef"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    77
fun get_multiset_setup thy = MultisetSetup.get thy
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    78
  |> the_default (Multiset
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    79
{ msetT = undef, mk_mset=undef,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    80
  mset_regroup_conv=undef, mset_member_tac = undef,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    81
  mset_nonempty_tac = undef, mset_pwleq_tac = undef,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    82
  set_of_simps = [],reduction_pair = refl,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    83
  smsI'=refl, wmsI2''=refl, wmsI1=refl })
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    84
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    85
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
    86
  | order_rpair msrp MS  = msrp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    87
  | order_rpair _ MIN = @{thm min_rpair_set}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    88
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    89
fun ord_intros_max true =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    90
    (@{thm smax_emptyI}, @{thm smax_insertI})
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    91
  | ord_intros_max false =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    92
    (@{thm wmax_emptyI}, @{thm wmax_insertI})
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    93
fun ord_intros_min true =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    94
    (@{thm smin_emptyI}, @{thm smin_insertI})
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    95
  | ord_intros_min false =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
    96
    (@{thm wmin_emptyI}, @{thm wmin_insertI})
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 gen_probl D cs =
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 n = Termination.get_num_points D
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   101
    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
   102
    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
   103
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   104
    fun mk_graph c =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   105
      let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   106
        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
   107
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   108
        fun add_edge i j =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   109
          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
   110
           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
   111
            | SOME (Termination.LessEq _) => cons (i, GEQ, j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   112
            | _ => I
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   113
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   114
        val edges =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   115
          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
   116
      in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   117
        G (p, q, edges)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   118
      end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   119
  in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   120
    GP (map arity (0 upto n - 1), map mk_graph cs)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   121
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   122
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   123
(* General reduction pair application *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   124
fun rem_inv_img ctxt =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   125
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   126
    val unfold_tac = LocalDefs.unfold_tac ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   127
  in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   128
    rtac @{thm subsetI} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   129
    THEN etac @{thm CollectE} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   130
    THEN REPEAT (etac @{thm exE} 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   131
    THEN unfold_tac @{thms inv_image_def}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   132
    THEN rtac @{thm CollectI} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   133
    THEN etac @{thm conjE} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   134
    THEN etac @{thm ssubst} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   135
    THEN unfold_tac (@{thms split_conv} @ @{thms triv_forall_equality}
29183
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 29125
diff changeset
   136
                     @ @{thms sum.cases})
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   137
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   138
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   139
(* Sets *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   140
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   141
val setT = HOLogic.mk_setT
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
fun mk_set T [] = Const (@{const_name "{}"}, setT T)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   144
  | mk_set T (x :: xs) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   145
      Const (@{const_name insert}, T --> setT T --> setT T) $
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   146
            x $ mk_set T xs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   147
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   148
fun set_member_tac m i =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   149
  if m = 0 then rtac @{thm insertI1} i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   150
  else rtac @{thm insertI2} i THEN set_member_tac (m - 1) i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   151
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   152
val set_nonempty_tac = rtac @{thm insert_not_empty}
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 set_finite_tac i =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   155
  rtac @{thm finite.emptyI} i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   156
  ORELSE (rtac @{thm finite.insertI} i THEN (fn st => set_finite_tac i st))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   157
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   158
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   159
(* Reconstruction *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   160
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   161
fun reconstruct_tac ctxt D cs (gp as GP (_, gs)) certificate =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   162
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   163
    val thy = ProofContext.theory_of ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   164
    val Multiset
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   165
          { msetT, mk_mset,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   166
            mset_regroup_conv, mset_member_tac,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   167
            mset_nonempty_tac, mset_pwleq_tac, set_of_simps,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   168
            smsI', wmsI2'', wmsI1, reduction_pair=ms_rp } 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   169
        = get_multiset_setup thy
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   170
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   171
    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
   172
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   173
    fun get_desc_thm cidx m1 m2 bStrict =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   174
      case Termination.get_descent D (nth cs cidx) m1 m2
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   175
       of SOME (Termination.Less thm) =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   176
          if bStrict then thm
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   177
          else (thm COMP (Thm.lift_rule (cprop_of thm) @{thm less_imp_le}))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   178
        | SOME (Termination.LessEq (thm, _))  =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   179
          if not bStrict then thm
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   180
          else sys_error "get_desc_thm"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   181
        | _ => sys_error "get_desc_thm"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   182
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   183
    val (label, lev, sl, covering) = certificate
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   184
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   185
    fun prove_lev strict g =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   186
      let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   187
        val G (p, q, el) = nth gs g
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   188
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   189
        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
   190
          let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   191
            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
   192
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   193
            val stored_thm =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   194
              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
   195
                             (not tag_flag)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   196
              |> Conv.fconv_rule (Thm.beta_conversion true)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   197
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   198
            val rule = if strict
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   199
              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
   200
              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
   201
          in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   202
            rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   203
            THEN (if tag_flag then arith_tac ctxt 1 else all_tac)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   204
          end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   205
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   206
        fun steps_tac MAX strict lq lp =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   207
          let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   208
            val (empty, step) = ord_intros_max strict
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   209
          in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   210
            if length lq = 0
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   211
            then rtac empty 1 THEN set_finite_tac 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   212
                 THEN (if strict then set_nonempty_tac 1 else all_tac)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   213
            else
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   214
              let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   215
                val (j, b) :: rest = lq
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   216
                val (i, a) = the (covering g strict j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   217
                fun choose xs = set_member_tac (Library.find_index (curry op = (i, a)) xs) 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   218
                val solve_tac = choose lp THEN less_proof strict (j, b) (i, a)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   219
              in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   220
                rtac step 1 THEN solve_tac THEN steps_tac MAX strict rest lp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   221
              end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   222
          end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   223
          | steps_tac MIN strict lq lp =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   224
          let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   225
            val (empty, step) = ord_intros_min strict
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   226
          in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   227
            if length lp = 0
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   228
            then rtac empty 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   229
                 THEN (if strict then set_nonempty_tac 1 else all_tac)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   230
            else
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   231
              let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   232
                val (i, a) :: rest = lp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   233
                val (j, b) = the (covering g strict i)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   234
                fun choose xs = set_member_tac (Library.find_index (curry op = (j, b)) xs) 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   235
                val solve_tac = choose lq THEN less_proof strict (j, b) (i, a)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   236
              in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   237
                rtac step 1 THEN solve_tac THEN steps_tac MIN strict lq rest
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   238
              end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   239
          end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   240
          | steps_tac MS strict lq lp =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   241
          let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   242
            fun get_str_cover (j, b) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   243
              if is_some (covering g true j) then SOME (j, b) else NONE
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   244
            fun get_wk_cover (j, b) = the (covering g false j)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   245
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   246
            val qs = lq \\ map_filter get_str_cover lq
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   247
            val ps = map get_wk_cover qs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   248
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   249
            fun indices xs ys = map (fn y => Library.find_index (curry op = y) xs) ys
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   250
            val iqs = indices lq qs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   251
            val ips = indices lp ps
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   252
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   253
            local open Conv in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   254
            fun t_conv a C =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   255
              params_conv ~1 (K ((concl_conv ~1 o arg_conv o arg1_conv o a) C)) ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   256
            val goal_rewrite =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   257
                t_conv arg1_conv (mset_regroup_conv iqs)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   258
                then_conv t_conv arg_conv (mset_regroup_conv ips)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   259
            end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   260
          in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   261
            CONVERSION goal_rewrite 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   262
            THEN (if strict then rtac smsI' 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   263
                  else if qs = lq then rtac wmsI2'' 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   264
                  else rtac wmsI1 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   265
            THEN mset_pwleq_tac 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   266
            THEN EVERY (map2 (less_proof false) qs ps)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   267
            THEN (if strict orelse qs <> lq
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   268
                  then LocalDefs.unfold_tac ctxt set_of_simps
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   269
                       THEN steps_tac MAX true (lq \\ qs) (lp \\ ps)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   270
                  else all_tac)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   271
          end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   272
      in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   273
        rem_inv_img ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   274
        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
   275
      end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   276
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   277
    val (mk_set, setT) = if label = MS then (mk_mset, msetT) else (mk_set, setT)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   278
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   279
    fun tag_pair p (i, tag) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   280
      HOLogic.pair_const natT natT $
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   281
        (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
   282
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   283
    fun pt_lev (p, lm) = Abs ("x", Termination.get_types D p,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   284
                           mk_set nat_pairT (map (tag_pair p) lm))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   285
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   286
    val level_mapping =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   287
      map_index pt_lev lev
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   288
        |> Termination.mk_sumcases D (setT nat_pairT)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   289
        |> cterm_of thy
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   290
    in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   291
      PROFILE "Proof Reconstruction"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   292
        (CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv sl))) 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   293
         THEN (rtac @{thm reduction_pair_lemma} 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   294
         THEN (rtac @{thm rp_inv_image_rp} 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   295
         THEN (rtac (order_rpair ms_rp label) 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   296
         THEN PRIMITIVE (instantiate' [] [SOME level_mapping])
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   297
         THEN unfold_tac @{thms rp_inv_image_def} (simpset_of thy)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   298
         THEN LocalDefs.unfold_tac ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   299
           (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv})
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   300
         THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}]))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   301
         THEN EVERY (map (prove_lev true) sl)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   302
         THEN EVERY (map (prove_lev false) ((0 upto length cs - 1) \\ sl)))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   303
    end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   304
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   305
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   306
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   307
local open Termination in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   308
fun print_cell (SOME (Less _)) = "<"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   309
  | print_cell (SOME (LessEq _)) = "\<le>"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   310
  | print_cell (SOME (None _)) = "-"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   311
  | print_cell (SOME (False _)) = "-"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   312
  | print_cell (NONE) = "?"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   313
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   314
fun print_error ctxt D = CALLS (fn (cs, i) =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   315
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   316
    val np = get_num_points D
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   317
    val ms = map (get_measures D) (0 upto np - 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   318
    val tys = map (get_types D) (0 upto np - 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   319
    fun index xs = (1 upto length xs) ~~ xs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   320
    fun outp s t f xs = map (fn (x, y) => s ^ Int.toString x ^ t ^ f y ^ "\n") xs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   321
    val ims = index (map index ms)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   322
    val _ = Output.tracing (concat (outp "fn #" ":\n" (concat o outp "\tmeasure #" ": " (Syntax.string_of_term ctxt)) ims))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   323
    fun print_call (k, c) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   324
      let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   325
        val (_, p, _, q, _, _) = dest_call D c
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   326
        val _ = Output.tracing ("call table for call #" ^ Int.toString k ^ ": fn " ^ 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   327
                                Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   328
        val caller_ms = nth ms p
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   329
        val callee_ms = nth ms q
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   330
        val entries = map (fn x => map (pair x) (callee_ms)) (caller_ms)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   331
        fun print_ln (i : int, l) = concat (Int.toString i :: "   " :: map (enclose " " " " o print_cell o (uncurry (get_descent D c))) l)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   332
        val _ = Output.tracing (concat (Int.toString (p + 1) ^ "|" ^ Int.toString (q + 1) ^ 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   333
                                        " " :: map (enclose " " " " o Int.toString) (1 upto length callee_ms)) ^ "\n" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   334
                                ^ cat_lines (map print_ln ((1 upto (length entries)) ~~ entries)))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   335
      in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   336
        true
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   337
      end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   338
    fun list_call (k, c) =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   339
      let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   340
        val (_, p, _, q, _, _) = dest_call D c
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   341
        val _ = Output.tracing ("call #" ^ (Int.toString k) ^ ": fn " ^
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   342
                                Int.toString (p + 1) ^ " ~> fn " ^ Int.toString (q + 1) ^ "\n" ^ 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   343
                                (Syntax.string_of_term ctxt c))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   344
      in true end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   345
    val _ = forall list_call ((1 upto length cs) ~~ cs)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   346
    val _ = forall print_call ((1 upto length cs) ~~ cs)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   347
  in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   348
    all_tac
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   349
  end)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   350
end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   351
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   352
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   353
fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   354
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   355
    val gp = gen_probl D cs
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   356
(*    val _ = TRACE ("SCNP instance: " ^ makestring gp)*)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   357
    val certificate = generate_certificate use_tags orders gp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   358
(*    val _ = TRACE ("Certificate: " ^ makestring certificate)*)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   359
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   360
    val ms_configured = is_some (MultisetSetup.get (ProofContext.theory_of ctxt))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   361
    in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   362
    case certificate
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   363
     of NONE => err_cont D i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   364
      | SOME cert =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   365
        if not ms_configured andalso #1 cert = MS
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   366
        then err_cont D i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   367
        else SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   368
             THEN (rtac @{thm wf_empty} i ORELSE cont D i)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   369
  end)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   370
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   371
fun decomp_scnp_tac orders autom_tac ctxt err_cont =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   372
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   373
    open Termination
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   374
    val derive_diag = Descent.derive_diag ctxt autom_tac
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   375
    val derive_all = Descent.derive_all ctxt autom_tac
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   376
    val decompose = Decompose.decompose_tac ctxt autom_tac
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   377
    val scnp_no_tags = single_scnp_tac false orders ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   378
    val scnp_full = single_scnp_tac true orders ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   379
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   380
    fun first_round c e =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   381
        derive_diag (REPEAT scnp_no_tags c e)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   382
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   383
    val second_round =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   384
        REPEAT (fn c => fn e => decompose (scnp_no_tags c c) e)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   385
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   386
    val third_round =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   387
        derive_all oo
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   388
        REPEAT (fn c => fn e =>
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   389
          scnp_full (decompose c c) e)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   390
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   391
    fun Then s1 s2 c e = s1 (s2 c c) (s2 c e)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   392
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   393
    val strategy = Then (Then first_round second_round) third_round
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   394
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   395
  in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   396
    TERMINATION ctxt (strategy err_cont err_cont)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   397
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   398
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   399
fun decomp_scnp orders ctxt =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   400
  let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   401
    val extra_simps = FundefCommon.TerminationSimps.get ctxt
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   402
    val autom_tac = auto_tac (local_clasimpset_of ctxt addsimps2 extra_simps)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   403
  in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   404
    Method.SIMPLE_METHOD
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   405
      (TRY (FundefCommon.apply_termination_rule ctxt 1)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   406
       THEN TRY Termination.wf_union_tac
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   407
       THEN
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   408
         (rtac @{thm wf_empty} 1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   409
          ORELSE decomp_scnp_tac orders autom_tac ctxt (print_error ctxt) 1))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   410
  end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   411
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   412
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   413
(* Method setup *)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   414
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   415
val orders =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   416
  (Scan.repeat1
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   417
    ((Args.$$$ "max" >> K MAX) ||
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   418
     (Args.$$$ "min" >> K MIN) ||
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   419
     (Args.$$$ "ms" >> K MS))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   420
  || Scan.succeed [MAX, MS, MIN])
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   421
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   422
val setup = Method.add_method
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   423
  ("sizechange", Method.sectioned_args (Scan.lift orders) clasimp_modifiers decomp_scnp,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   424
   "termination prover with graph decomposition and the NP subset of size change termination")
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   425
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
diff changeset
   426
end