| author | chaieb | 
| Fri, 27 Mar 2009 17:35:21 +0000 | |
| changeset 30748 | fe67d729a61c | 
| parent 30715 | e23e15f52d42 | 
| child 32149 | ef59550a55d3 | 
| permissions | -rw-r--r-- | 
| 28308 | 1  | 
(* Title: HOL/SizeChange/sct.ML  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
2  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
4  | 
Tactics for size change termination.  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
6  | 
signature SCT =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
7  | 
sig  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
8  | 
val abs_rel_tac : tactic  | 
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26187 
diff
changeset
 | 
9  | 
val mk_call_graph : Proof.context -> tactic  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
10  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
12  | 
structure Sct : SCT =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
13  | 
struct  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
15  | 
fun matrix [] ys = []  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
16  | 
| matrix (x::xs) ys = map (pair x) ys :: matrix xs ys  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
18  | 
fun map_matrix f xss = map (map f) xss  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
20  | 
val scgT = @{typ "nat scg"}
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
21  | 
val acgT = @{typ "nat acg"}
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
23  | 
fun edgeT nT eT = HOLogic.mk_prodT (nT, HOLogic.mk_prodT (eT, nT))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
24  | 
fun graphT nT eT = Type ("Graphs.graph", [nT, eT])
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
26  | 
fun graph_const nT eT = Const ("Graphs.graph.Graph", HOLogic.mk_setT (edgeT nT eT) --> graphT nT eT)
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
28  | 
val stepP_const = "Interpretation.stepP"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
29  | 
val stepP_def = thm "Interpretation.stepP.simps"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
31  | 
fun mk_stepP RD1 RD2 M1 M2 Rel =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
32  | 
let val RDT = fastype_of RD1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
33  | 
val MT = fastype_of M1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
34  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
35  | 
Const (stepP_const, RDT --> RDT --> MT --> MT --> (fastype_of Rel) --> HOLogic.boolT)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
36  | 
$ RD1 $ RD2 $ M1 $ M2 $ Rel  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
37  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
39  | 
val no_stepI = thm "Interpretation.no_stepI"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
41  | 
val approx_const = "Interpretation.approx"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
42  | 
val approx_empty = thm "Interpretation.approx_empty"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
43  | 
val approx_less = thm "Interpretation.approx_less"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
44  | 
val approx_leq = thm "Interpretation.approx_leq"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
46  | 
fun mk_approx G RD1 RD2 Ms1 Ms2 =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
47  | 
let val RDT = fastype_of RD1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
48  | 
val MsT = fastype_of Ms1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
49  | 
in Const (approx_const, scgT --> RDT --> RDT --> MsT --> MsT --> HOLogic.boolT) $ G $ RD1 $ RD2 $ Ms1 $ Ms2 end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
51  | 
val sound_int_const = "Interpretation.sound_int"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
52  | 
val sound_int_def = thm "Interpretation.sound_int_def"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
53  | 
fun mk_sound_int A RDs M =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
54  | 
let val RDsT = fastype_of RDs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
55  | 
val MT = fastype_of M  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
56  | 
in Const (sound_int_const, acgT --> RDsT --> MT --> HOLogic.boolT) $ A $ RDs $ M end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
59  | 
val nth_const = "List.nth"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
60  | 
fun mk_nth xs =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
61  | 
let val lT as Type (_, [T]) = fastype_of xs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
62  | 
in Const (nth_const, lT --> HOLogic.natT --> T) $ xs end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
65  | 
val has_edge_simps = [thm "Graphs.has_edge_def", thm "Graphs.dest_graph.simps"]  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
67  | 
val all_less_zero = thm "Interpretation.all_less_zero"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
68  | 
val all_less_Suc = thm "Interpretation.all_less_Suc"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
70  | 
(* --> Library? *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
71  | 
fun del_index n [] = []  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
72  | 
| del_index n (x :: xs) =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
73  | 
if n>0 then x :: del_index (n - 1) xs else xs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
74  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
75  | 
(* Lists as finite multisets *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
77  | 
fun remove1 eq x [] = []  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
78  | 
| remove1 eq x (y :: ys) = if eq (x, y) then ys else y :: remove1 eq x ys  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
80  | 
fun multi_union eq [] ys = ys  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
81  | 
| multi_union eq (x::xs) ys = x :: multi_union eq xs (remove1 eq x ys)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
82  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
83  | 
fun dest_ex (Const ("Ex", _) $ Abs (a as (_,T,_))) =
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
84  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
85  | 
val (n, body) = Term.dest_abs a  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
86  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
87  | 
(Free (n, T), body)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
88  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
89  | 
| dest_ex _ = raise Match  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
91  | 
fun dest_all_ex (t as (Const ("Ex",_) $ _)) =
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
92  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
93  | 
val (v,b) = dest_ex t  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
94  | 
val (vs, b') = dest_all_ex b  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
95  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
96  | 
(v :: vs, b')  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
97  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
98  | 
| dest_all_ex t = ([],t)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
100  | 
fun dist_vars [] vs = (null vs orelse error "dist_vars"; [])  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
101  | 
| dist_vars (T::Ts) vs =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
102  | 
case find_index (fn v => fastype_of v = T) vs of  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
103  | 
      ~1 => Free ("", T) :: dist_vars Ts vs
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
104  | 
| i => (nth vs i) :: dist_vars Ts (del_index i vs)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
106  | 
fun dest_case rebind t =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
107  | 
let  | 
| 
30715
 
e23e15f52d42
avoid mixing of left/right associative infixes, to make it work with experimental Poly/ML 5.3 branch;
 
wenzelm 
parents: 
30607 
diff
changeset
 | 
108  | 
val ((_ $ _ $ rhs) :: (_ $ _ $ match) :: guards) = HOLogic.dest_conj t  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
109  | 
val guard = case guards of [] => HOLogic.true_const | gs => foldr1 HOLogic.mk_conj gs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
110  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
111  | 
foldr1 HOLogic.mk_prod [rebind guard, rebind rhs, rebind match]  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
112  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
114  | 
fun bind_many [] = I  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
115  | 
| bind_many vs = FundefLib.tupled_lambda (foldr1 HOLogic.mk_prod vs)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
116  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
117  | 
(* Builds relation descriptions from a relation definition *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
118  | 
fun mk_reldescs (Abs a) =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
119  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
120  | 
val (_, Abs a') = Term.dest_abs a  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
121  | 
val (_, b) = Term.dest_abs a'  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
122  | 
val cases = HOLogic.dest_disj b  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
123  | 
val (vss, bs) = split_list (map dest_all_ex cases)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
124  | 
val unionTs = fold (multi_union (op =)) (map (map fastype_of) vss) []  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
125  | 
val rebind = map (bind_many o dist_vars unionTs) vss  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
126  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
127  | 
val RDs = map2 dest_case rebind bs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
128  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
129  | 
HOLogic.mk_list (fastype_of (hd RDs)) RDs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
130  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
132  | 
fun abs_rel_tac (st : thm) =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
133  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
134  | 
val thy = theory_of_thm st  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
135  | 
val (def, rd) = HOLogic.dest_eq (HOLogic.dest_Trueprop (hd (prems_of st)))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
136  | 
val RDs = cterm_of thy (mk_reldescs def)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
137  | 
val rdvar = Var (the_single (Term.add_vars rd [])) |> cterm_of thy  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
138  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
139  | 
Seq.single (cterm_instantiate [(rdvar, RDs)] st)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
140  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
141  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
142  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
145  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
147  | 
(* very primitive *)  | 
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26187 
diff
changeset
 | 
148  | 
fun measures_of ctxt RD =  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
149  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
150  | 
val domT = range_type (fastype_of (fst (HOLogic.dest_prod (snd (HOLogic.dest_prod RD)))))  | 
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26187 
diff
changeset
 | 
151  | 
val measures = MeasureFunctions.get_measure_functions ctxt domT  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
152  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
153  | 
measures  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
154  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
155  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
156  | 
val mk_number = HOLogic.mk_nat  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
157  | 
val dest_number = HOLogic.dest_nat  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
159  | 
fun nums_to i = map mk_number (0 upto (i - 1))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
161  | 
val nth_simps = [thm "List.nth_Cons_0", thm "List.nth_Cons_Suc"]  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
162  | 
val nth_ss = (HOL_basic_ss addsimps nth_simps)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
163  | 
val simp_nth_tac = simp_tac nth_ss  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
166  | 
fun tabulate_tlist thy l =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
167  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
168  | 
val n = length (HOLogic.dest_list l)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
169  | 
val table = Inttab.make (map (fn i => (i, Simplifier.rewrite nth_ss (cterm_of thy (mk_nth l $ mk_number i)))) (0 upto n - 1))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
170  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
171  | 
the o Inttab.lookup table  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
172  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
174  | 
val get_elem = snd o Logic.dest_equals o prop_of  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
176  | 
fun inst_nums thy i j (t:thm) =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
177  | 
instantiate' [] [NONE, NONE, NONE, SOME (cterm_of thy (mk_number i)), NONE, SOME (cterm_of thy (mk_number j))] t  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
178  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
179  | 
datatype call_fact =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
180  | 
NoStep of thm  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
181  | 
| Graph of (term * thm)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
183  | 
fun rand (_ $ t) = t  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
184  | 
|
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
185  | 
fun setup_probe_goal ctxt domT Dtab Mtab (i, j) =  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
186  | 
let  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
187  | 
val css = local_clasimpset_of ctxt  | 
| 
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
188  | 
val thy = ProofContext.theory_of ctxt  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
189  | 
val RD1 = get_elem (Dtab i)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
190  | 
val RD2 = get_elem (Dtab j)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
191  | 
val Ms1 = get_elem (Mtab i)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
192  | 
val Ms2 = get_elem (Mtab j)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
193  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
194  | 
val Mst1 = HOLogic.dest_list (rand Ms1)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
195  | 
val Mst2 = HOLogic.dest_list (rand Ms2)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
196  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
197  | 
      val mvar1 = Free ("sctmfv1", domT --> HOLogic.natT)
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
198  | 
      val mvar2 = Free ("sctmfv2", domT --> HOLogic.natT)
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
199  | 
      val relvar = Free ("sctmfrel", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT)
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
200  | 
val N = length Mst1 and M = length Mst2  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
201  | 
val saved_state = HOLogic.mk_Trueprop (mk_stepP RD1 RD2 mvar1 mvar2 relvar)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
202  | 
|> cterm_of thy  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
203  | 
|> Goal.init  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
204  | 
|> auto_tac css |> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
205  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
206  | 
val no_step = saved_state  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
207  | 
|> forall_intr (cterm_of thy relvar)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
208  | 
                      |> forall_elim (cterm_of thy (Abs ("", HOLogic.natT, Abs ("", HOLogic.natT, HOLogic.false_const))))
 | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
209  | 
|> auto_tac css |> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
211  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
212  | 
if Thm.no_prems no_step  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
213  | 
then NoStep (Goal.finish no_step RS no_stepI)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
214  | 
else  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
215  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
216  | 
fun set_m1 i =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
217  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
218  | 
val M1 = nth Mst1 i  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
219  | 
val with_m1 = saved_state  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
220  | 
|> forall_intr (cterm_of thy mvar1)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
221  | 
|> forall_elim (cterm_of thy M1)  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
222  | 
|> auto_tac css |> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
223  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
224  | 
fun set_m2 j =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
225  | 
let  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
226  | 
val M2 = nth Mst2 j  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
227  | 
val with_m2 = with_m1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
228  | 
|> forall_intr (cterm_of thy mvar2)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
229  | 
|> forall_elim (cterm_of thy M2)  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
230  | 
|> auto_tac css |> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
232  | 
val decr = forall_intr (cterm_of thy relvar)  | 
| 26187 | 233  | 
                                   #> forall_elim (cterm_of thy @{const HOL.less(nat)})
 | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
234  | 
#> auto_tac css #> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
236  | 
val decreq = forall_intr (cterm_of thy relvar)  | 
| 26187 | 237  | 
                                     #> forall_elim (cterm_of thy @{const HOL.less_eq(nat)})
 | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
238  | 
#> auto_tac css #> Seq.hd  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
239  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
240  | 
val thm1 = decr with_m2  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
241  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
242  | 
if Thm.no_prems thm1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
243  | 
then ((rtac (inst_nums thy i j approx_less) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish thm1) 1))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
244  | 
else let val thm2 = decreq with_m2 in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
245  | 
if Thm.no_prems thm2  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
246  | 
then ((rtac (inst_nums thy i j approx_leq) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish thm2) 1))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
247  | 
else all_tac end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
248  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
249  | 
in set_m2 end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
250  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
251  | 
          val goal = HOLogic.mk_Trueprop (mk_approx (Var (("G", 0), scgT)) RD1 RD2 Ms1 Ms2)
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
252  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
253  | 
val tac = (EVERY (map (fn n => EVERY (map (set_m1 n) (0 upto M - 1))) (0 upto N - 1)))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
254  | 
THEN (rtac approx_empty 1)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
255  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
256  | 
val approx_thm = goal  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
257  | 
|> cterm_of thy  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
258  | 
|> Goal.init  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
259  | 
|> tac |> Seq.hd  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
260  | 
|> Goal.finish  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
262  | 
val _ $ (_ $ G $ _ $ _ $ _ $ _) = prop_of approx_thm  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
263  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
264  | 
Graph (G, approx_thm)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
265  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
266  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
267  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
268  | 
fun mk_edge m G n = HOLogic.mk_prod (m, HOLogic.mk_prod (G, n))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
269  | 
|
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
270  | 
fun in_graph_tac ctxt =  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
271  | 
simp_tac (HOL_basic_ss addsimps has_edge_simps) 1  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
272  | 
THEN (simp_tac (local_simpset_of ctxt) 1) (* FIXME reduce simpset *)  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
273  | 
|
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
274  | 
fun approx_tac _ (NoStep thm) = rtac disjI1 1 THEN rtac thm 1  | 
| 
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
275  | 
| approx_tac ctxt (Graph (G, thm)) =  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
276  | 
rtac disjI2 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
277  | 
THEN rtac exI 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
278  | 
THEN rtac conjI 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
279  | 
THEN rtac thm 2  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
280  | 
THEN (in_graph_tac ctxt)  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
281  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
282  | 
fun all_less_tac [] = rtac all_less_zero 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
283  | 
| all_less_tac (t :: ts) = rtac all_less_Suc 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
284  | 
THEN simp_nth_tac 1  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
285  | 
THEN t  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
286  | 
THEN all_less_tac ts  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
289  | 
fun mk_length l = HOLogic.size_const (fastype_of l) $ l;  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
290  | 
val length_simps = thms "Interpretation.length_simps"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
291  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
292  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
293  | 
|
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26187 
diff
changeset
 | 
294  | 
fun mk_call_graph ctxt (st : thm) =  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
295  | 
let  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
296  | 
val thy = ProofContext.theory_of ctxt  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
297  | 
val _ $ _ $ RDlist $ _ = HOLogic.dest_Trueprop (hd (prems_of st))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
298  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
299  | 
val RDs = HOLogic.dest_list RDlist  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
300  | 
val n = length RDs  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
301  | 
|
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26187 
diff
changeset
 | 
302  | 
val Mss = map (measures_of ctxt) RDs  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
303  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
304  | 
val domT = domain_type (fastype_of (hd (hd Mss)))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
305  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
306  | 
val mfuns = map (fn Ms => mk_nth (HOLogic.mk_list (fastype_of (hd Ms)) Ms)) Mss  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
307  | 
|> (fn l => HOLogic.mk_list (fastype_of (hd l)) l)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
308  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
309  | 
val Dtab = tabulate_tlist thy RDlist  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
310  | 
val Mtab = tabulate_tlist thy mfuns  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
312  | 
val len_simp = Simplifier.rewrite (HOL_basic_ss addsimps length_simps) (cterm_of thy (mk_length RDlist))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
313  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
314  | 
val mlens = map length Mss  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
315  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
316  | 
val indices = (n - 1 downto 0)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
317  | 
val pairs = matrix indices indices  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
318  | 
val parts = map_matrix (fn (n,m) =>  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
319  | 
(timeap_msg (string_of_int n ^ "," ^ string_of_int m)  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
320  | 
(setup_probe_goal ctxt domT Dtab Mtab) (n,m))) pairs  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
321  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
322  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
323  | 
      val s = fold_index (fn (i, cs) => fold_index (fn (j, Graph (G, _)) => prefix ("(" ^ string_of_int i ^ "," ^ string_of_int j ^ "): " ^
 | 
| 26953 | 324  | 
Syntax.string_of_term ctxt G ^ ",\n")  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
325  | 
| _ => I) cs) parts ""  | 
| 26941 | 326  | 
val _ = warning s  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
327  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
328  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
329  | 
val ACG = map_filter (fn (Graph (G, _),(m, n)) => SOME (mk_edge (mk_number m) G (mk_number n)) | _ => NONE) (flat parts ~~ flat pairs)  | 
| 30450 | 330  | 
|> HOLogic.mk_set (edgeT HOLogic.natT scgT)  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
331  | 
|> curry op $ (graph_const HOLogic.natT scgT)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
332  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
333  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
334  | 
val sound_int_goal = HOLogic.mk_Trueprop (mk_sound_int ACG RDlist mfuns)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
335  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
336  | 
val tac =  | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
337  | 
unfold_tac [sound_int_def, len_simp] (local_simpset_of ctxt)  | 
| 
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
30450 
diff
changeset
 | 
338  | 
THEN all_less_tac (map (all_less_tac o map (approx_tac ctxt)) parts)  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
339  | 
in  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
340  | 
tac (instantiate' [] [SOME (cterm_of thy ACG), SOME (cterm_of thy mfuns)] st)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
341  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
342  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
343  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
344  | 
end  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
345  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
347  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
348  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
349  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
350  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
351  |