| author | wenzelm | 
| Wed, 28 Oct 2009 22:26:00 +0100 | |
| changeset 33293 | 4645818f0fbd | 
| parent 33099 | b8cdd3d73022 | 
| child 33351 | 37ec56ac3fd4 | 
| child 33569 | 1ebb8b7b9f6a | 
| permissions | -rw-r--r-- | 
| 31775 | 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  | 
|
| 
 
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  | 
|
| 29877 | 11  | 
val sizechange_tac : Proof.context -> tactic -> tactic  | 
12  | 
||
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
13  | 
val decomp_scnp : ScnpSolve.label list -> Proof.context -> Proof.method  | 
| 
29125
 
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  | 
val setup : theory -> theory  | 
| 
 
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  | 
datatype multiset_setup =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
18  | 
Multiset of  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
19  | 
    {
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
20  | 
msetT : typ -> typ,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
21  | 
mk_mset : typ -> term list -> term,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
22  | 
mset_regroup_conv : int list -> conv,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
23  | 
mset_member_tac : int -> int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
24  | 
mset_nonempty_tac : int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
25  | 
mset_pwleq_tac : int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
26  | 
set_of_simps : thm list,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
27  | 
smsI' : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
28  | 
wmsI2'' : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
29  | 
wmsI1 : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
30  | 
reduction_pair : thm  | 
| 
 
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  | 
|
| 
 
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  | 
val multiset_setup : multiset_setup -> theory -> theory  | 
| 
 
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  | 
end  | 
| 
 
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  | 
structure ScnpReconstruct : SCNP_RECONSTRUCT =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
39  | 
struct  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
40  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
41  | 
val PROFILE = Function_Common.PROFILE  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
42  | 
fun TRACE x = if ! Function_Common.profile then tracing x else ()  | 
| 
29125
 
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  | 
open ScnpSolve  | 
| 
 
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  | 
val natT = HOLogic.natT  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
47  | 
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
 | 
48  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
49  | 
(* Theory dependencies *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
51  | 
datatype multiset_setup =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
52  | 
Multiset of  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
53  | 
  {
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
54  | 
msetT : typ -> typ,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
55  | 
mk_mset : typ -> term list -> term,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
56  | 
mset_regroup_conv : int list -> conv,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
57  | 
mset_member_tac : int -> int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
58  | 
mset_nonempty_tac : int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
59  | 
mset_pwleq_tac : int -> tactic,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
60  | 
set_of_simps : thm list,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
61  | 
smsI' : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
62  | 
wmsI2'' : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
63  | 
wmsI1 : thm,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
64  | 
reduction_pair : thm  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
65  | 
}  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
66  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
67  | 
structure Multiset_Setup = TheoryDataFun  | 
| 
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  | 
type T = multiset_setup option  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
70  | 
val empty = NONE  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
71  | 
val copy = I;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
72  | 
val extend = I;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
73  | 
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
 | 
74  | 
)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
75  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
76  | 
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
 | 
77  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
78  | 
fun undef x = error "undef"  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
79  | 
fun get_multiset_setup thy = Multiset_Setup.get thy  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
80  | 
|> the_default (Multiset  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
81  | 
{ msetT = undef, mk_mset=undef,
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
82  | 
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
 | 
83  | 
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
 | 
84  | 
set_of_simps = [],reduction_pair = refl,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
85  | 
smsI'=refl, wmsI2''=refl, wmsI1=refl })  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
87  | 
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
 | 
88  | 
| order_rpair msrp MS = msrp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
89  | 
  | order_rpair _ MIN = @{thm min_rpair_set}
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
91  | 
fun ord_intros_max true =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
92  | 
    (@{thm smax_emptyI}, @{thm smax_insertI})
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
93  | 
| ord_intros_max false =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
94  | 
    (@{thm wmax_emptyI}, @{thm wmax_insertI})
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
95  | 
fun ord_intros_min true =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
96  | 
    (@{thm smin_emptyI}, @{thm smin_insertI})
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
97  | 
| ord_intros_min false =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
98  | 
    (@{thm wmin_emptyI}, @{thm wmin_insertI})
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
100  | 
fun gen_probl D cs =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
101  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
102  | 
val n = Termination.get_num_points D  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
103  | 
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
 | 
104  | 
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
 | 
105  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
106  | 
fun mk_graph c =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
107  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
108  | 
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
 | 
109  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
110  | 
fun add_edge i j =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
111  | 
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
 | 
112  | 
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
 | 
113  | 
| SOME (Termination.LessEq _) => cons (i, GEQ, j)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
114  | 
| _ => I  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
116  | 
val edges =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
117  | 
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
 | 
118  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
119  | 
G (p, q, edges)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
120  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
121  | 
in  | 
| 33063 | 122  | 
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
 | 
123  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
125  | 
(* General reduction pair application *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
126  | 
fun rem_inv_img ctxt =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
127  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
128  | 
val unfold_tac = LocalDefs.unfold_tac ctxt  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
129  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
130  | 
    rtac @{thm subsetI} 1
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
131  | 
    THEN etac @{thm CollectE} 1
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
132  | 
    THEN REPEAT (etac @{thm exE} 1)
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
133  | 
    THEN unfold_tac @{thms inv_image_def}
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
134  | 
    THEN rtac @{thm CollectI} 1
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
135  | 
    THEN etac @{thm conjE} 1
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
136  | 
    THEN etac @{thm ssubst} 1
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
137  | 
    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
 | 
138  | 
                     @ @{thms sum.cases})
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
139  | 
end  | 
| 
 
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  | 
(* Sets *)  | 
| 
 
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  | 
val setT = HOLogic.mk_setT  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
145  | 
fun set_member_tac m i =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
146  | 
  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
 | 
147  | 
  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
 | 
148  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
149  | 
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
 | 
150  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
151  | 
fun set_finite_tac i =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
152  | 
  rtac @{thm finite.emptyI} i
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
153  | 
  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
 | 
154  | 
|
| 
 
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  | 
(* Reconstruction *)  | 
| 
 
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  | 
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
 | 
159  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
160  | 
val thy = ProofContext.theory_of ctxt  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
161  | 
val Multiset  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
162  | 
          { msetT, mk_mset,
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
163  | 
mset_regroup_conv, mset_member_tac,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
164  | 
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
 | 
165  | 
smsI', wmsI2'', wmsI1, reduction_pair=ms_rp }  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
166  | 
= get_multiset_setup thy  | 
| 
 
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 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
 | 
169  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
170  | 
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
 | 
171  | 
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
 | 
172  | 
of SOME (Termination.Less thm) =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
173  | 
if bStrict then thm  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
174  | 
          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
 | 
175  | 
| SOME (Termination.LessEq (thm, _)) =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
176  | 
if not bStrict then thm  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
177  | 
else sys_error "get_desc_thm"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
178  | 
| _ => sys_error "get_desc_thm"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
180  | 
val (label, lev, sl, covering) = certificate  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
182  | 
fun prove_lev strict g =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
183  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
184  | 
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
 | 
185  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
186  | 
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
 | 
187  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
188  | 
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
 | 
189  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
190  | 
val stored_thm =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
191  | 
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
 | 
192  | 
(not tag_flag)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
193  | 
|> Conv.fconv_rule (Thm.beta_conversion true)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
194  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
195  | 
val rule = if strict  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
196  | 
              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
 | 
197  | 
              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
 | 
198  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
199  | 
rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)  | 
| 
30686
 
47a32dd1b86e
moved generic arith_tac (formerly silent_arith_tac), verbose_arith_tac (formerly arith_tac) to Arith_Data; simple_arith-tac now named linear_arith_tac
 
haftmann 
parents: 
30607 
diff
changeset
 | 
200  | 
THEN (if tag_flag then Arith_Data.verbose_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
 | 
201  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
203  | 
fun steps_tac MAX strict lq lp =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
204  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
205  | 
val (empty, step) = ord_intros_max strict  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
206  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
207  | 
if length lq = 0  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
208  | 
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
 | 
209  | 
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
 | 
210  | 
else  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
211  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
212  | 
val (j, b) :: rest = lq  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
213  | 
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
 | 
214  | 
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
 | 
215  | 
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
 | 
216  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
217  | 
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
 | 
218  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
219  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
220  | 
| steps_tac MIN strict lq lp =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
221  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
222  | 
val (empty, step) = ord_intros_min strict  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
223  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
224  | 
if length lp = 0  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
225  | 
then rtac empty 1  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
226  | 
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
 | 
227  | 
else  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
228  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
229  | 
val (i, a) :: rest = lp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
230  | 
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
 | 
231  | 
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
 | 
232  | 
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
 | 
233  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
234  | 
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
 | 
235  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
236  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
237  | 
| steps_tac MS strict lq lp =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
238  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
239  | 
fun get_str_cover (j, b) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
240  | 
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
 | 
241  | 
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
 | 
242  | 
|
| 33040 | 243  | 
val qs = subtract (op =) (map_filter get_str_cover lq) lq  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
244  | 
val ps = map get_wk_cover qs  | 
| 
 
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  | 
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
 | 
247  | 
val iqs = indices lq qs  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
248  | 
val ips = indices lp ps  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
249  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
250  | 
local open Conv in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
251  | 
fun t_conv a C =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
252  | 
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
 | 
253  | 
val goal_rewrite =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
254  | 
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
 | 
255  | 
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
 | 
256  | 
end  | 
| 
 
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  | 
CONVERSION goal_rewrite 1  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
259  | 
THEN (if strict then rtac smsI' 1  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
260  | 
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
 | 
261  | 
else rtac wmsI1 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
262  | 
THEN mset_pwleq_tac 1  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
263  | 
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
 | 
264  | 
THEN (if strict orelse qs <> lq  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
265  | 
then LocalDefs.unfold_tac ctxt set_of_simps  | 
| 33040 | 266  | 
THEN steps_tac MAX true  | 
267  | 
(subtract (op =) qs lq) (subtract (op =) ps lp)  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
268  | 
else all_tac)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
269  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
270  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
271  | 
rem_inv_img ctxt  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
272  | 
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
 | 
273  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
274  | 
|
| 30450 | 275  | 
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
 | 
276  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
277  | 
fun tag_pair p (i, tag) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
278  | 
HOLogic.pair_const natT natT $  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
279  | 
(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
 | 
280  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
281  | 
    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
 | 
282  | 
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
 | 
283  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
284  | 
val level_mapping =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
285  | 
map_index pt_lev lev  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
286  | 
|> Termination.mk_sumcases D (setT nat_pairT)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
287  | 
|> cterm_of thy  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
288  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
289  | 
PROFILE "Proof Reconstruction"  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
290  | 
(CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv sl))) 1  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
291  | 
         THEN (rtac @{thm reduction_pair_lemma} 1)
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
292  | 
         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
 | 
293  | 
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
 | 
294  | 
THEN PRIMITIVE (instantiate' [] [SOME level_mapping])  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
31902 
diff
changeset
 | 
295  | 
         THEN unfold_tac @{thms rp_inv_image_def} (simpset_of ctxt)
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
296  | 
THEN LocalDefs.unfold_tac ctxt  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
297  | 
           (@{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
 | 
298  | 
         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
 | 
299  | 
THEN EVERY (map (prove_lev true) sl)  | 
| 33040 | 300  | 
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
 | 
301  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
302  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
303  | 
|
| 
 
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  | 
local open Termination in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
306  | 
fun print_cell (SOME (Less _)) = "<"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
307  | 
| print_cell (SOME (LessEq _)) = "\<le>"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
308  | 
| print_cell (SOME (None _)) = "-"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
309  | 
| print_cell (SOME (False _)) = "-"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
310  | 
| print_cell (NONE) = "?"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
312  | 
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
 | 
313  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
314  | 
val np = get_num_points D  | 
| 33063 | 315  | 
val ms = map_range (get_measures D) np  | 
316  | 
val tys = map_range (get_types D) np  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
317  | 
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
 | 
318  | 
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
 | 
319  | 
val ims = index (map index ms)  | 
| 32952 | 320  | 
val _ = tracing (implode (outp "fn #" ":\n" (implode o outp "\tmeasure #" ": " (Syntax.string_of_term ctxt)) ims))  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
321  | 
fun print_call (k, c) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
322  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
323  | 
val (_, p, _, q, _, _) = dest_call D c  | 
| 32950 | 324  | 
        val _ = tracing ("call table for call #" ^ Int.toString k ^ ": fn " ^ 
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
325  | 
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
 | 
326  | 
val caller_ms = nth ms p  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
327  | 
val callee_ms = nth ms q  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
328  | 
val entries = map (fn x => map (pair x) (callee_ms)) (caller_ms)  | 
| 32952 | 329  | 
fun print_ln (i : int, l) = implode (Int.toString i :: " " :: map (enclose " " " " o print_cell o (uncurry (get_descent D c))) l)  | 
330  | 
val _ = tracing (implode (Int.toString (p + 1) ^ "|" ^ Int.toString (q + 1) ^  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
331  | 
" " :: 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
 | 
332  | 
^ 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
 | 
333  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
334  | 
true  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
335  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
336  | 
fun list_call (k, c) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
337  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
338  | 
val (_, p, _, q, _, _) = dest_call D c  | 
| 32950 | 339  | 
        val _ = tracing ("call #" ^ (Int.toString k) ^ ": fn " ^
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
340  | 
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
 | 
341  | 
(Syntax.string_of_term ctxt c))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
342  | 
in true end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
343  | 
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
 | 
344  | 
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
 | 
345  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
346  | 
all_tac  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
347  | 
end)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
348  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
349  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
350  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
351  | 
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
 | 
352  | 
let  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
353  | 
val ms_configured = is_some (Multiset_Setup.get (ProofContext.theory_of ctxt))  | 
| 29877 | 354  | 
val orders' = if ms_configured then orders  | 
355  | 
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
 | 
356  | 
val gp = gen_probl D cs  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
357  | 
(*    val _ = TRACE ("SCNP instance: " ^ makestring gp)*)
 | 
| 29877 | 358  | 
val certificate = generate_certificate use_tags orders' gp  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
359  | 
(*    val _ = TRACE ("Certificate: " ^ makestring certificate)*)
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
360  | 
|
| 29877 | 361  | 
in  | 
| 
29125
 
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 =>  | 
| 29877 | 365  | 
SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i  | 
366  | 
          THEN (rtac @{thm wf_empty} i ORELSE cont D i)
 | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
367  | 
end)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
368  | 
|
| 29877 | 369  | 
fun gen_decomp_scnp_tac orders autom_tac ctxt err_cont =  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
370  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
371  | 
open Termination  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
372  | 
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
 | 
373  | 
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
 | 
374  | 
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
 | 
375  | 
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
 | 
376  | 
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
 | 
377  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
378  | 
fun first_round c e =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
379  | 
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
 | 
380  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
381  | 
val second_round =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
382  | 
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
 | 
383  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
384  | 
val third_round =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
385  | 
derive_all oo  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
386  | 
REPEAT (fn c => fn e =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
387  | 
scnp_full (decompose c c) e)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
388  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
389  | 
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
 | 
390  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
391  | 
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
 | 
392  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
393  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
394  | 
TERMINATION ctxt (strategy err_cont err_cont)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
395  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
396  | 
|
| 29877 | 397  | 
fun gen_sizechange_tac orders autom_tac ctxt err_cont =  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
398  | 
TRY (Function_Common.apply_termination_rule ctxt 1)  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30510 
diff
changeset
 | 
399  | 
THEN TRY (Termination.wf_union_tac ctxt)  | 
| 29877 | 400  | 
THEN  | 
401  | 
   (rtac @{thm wf_empty} 1
 | 
|
402  | 
ORELSE gen_decomp_scnp_tac orders autom_tac ctxt err_cont 1)  | 
|
403  | 
||
404  | 
fun sizechange_tac ctxt autom_tac =  | 
|
405  | 
gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt (K (K all_tac))  | 
|
406  | 
||
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
407  | 
fun decomp_scnp orders ctxt =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
408  | 
let  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33063 
diff
changeset
 | 
409  | 
val extra_simps = Function_Common.Termination_Simps.get ctxt  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
31902 
diff
changeset
 | 
410  | 
val autom_tac = auto_tac (clasimpset_of ctxt addsimps2 extra_simps)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
411  | 
in  | 
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
412  | 
SIMPLE_METHOD  | 
| 29877 | 413  | 
(gen_sizechange_tac orders autom_tac ctxt (print_error ctxt))  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
414  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
415  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
416  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
417  | 
(* Method setup *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
418  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
419  | 
val orders =  | 
| 31242 | 420  | 
Scan.repeat1  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
421  | 
((Args.$$$ "max" >> K MAX) ||  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
422  | 
(Args.$$$ "min" >> K MIN) ||  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
423  | 
(Args.$$$ "ms" >> K MS))  | 
| 31242 | 424  | 
|| Scan.succeed [MAX, MS, MIN]  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
425  | 
|
| 31242 | 426  | 
val setup = Method.setup @{binding sizechange}
 | 
427  | 
(Scan.lift orders --| Method.sections clasimp_modifiers >> decomp_scnp)  | 
|
428  | 
"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
 | 
429  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
430  | 
end  |