| author | krauss | 
| Thu, 18 Sep 2008 12:13:50 +0200 | |
| changeset 28287 | c86fa4e0aedb | 
| parent 27417 | 6cc897e2468a | 
| child 30952 | 7ab2716dd93b | 
| permissions | -rw-r--r-- | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Library/SCT_Interpretation.thy  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
3  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
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  | 
header {* Applying SCT to function definitions *}
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
8  | 
theory Interpretation  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
9  | 
imports Main Misc_Tools Criterion  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
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  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
13  | 
"idseq R s x = (s 0 = x \<and> (\<forall>i. R (s (Suc i)) (s i)))"  | 
| 
 
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  | 
lemma not_acc_smaller:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
16  | 
assumes notacc: "\<not> accp R x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
17  | 
shows "\<exists>y. R y x \<and> \<not> accp R y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
18  | 
proof (rule classical)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
19  | 
assume "\<not> ?thesis"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
20  | 
hence "\<And>y. R y x \<Longrightarrow> accp R y" by blast  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
21  | 
with accp.accI have "accp R x" .  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
22  | 
with notacc show ?thesis by contradiction  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
23  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
25  | 
lemma non_acc_has_idseq:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
26  | 
assumes "\<not> accp R x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
27  | 
shows "\<exists>s. idseq R s x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
28  | 
proof -  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
30  | 
have "\<exists>f. \<forall>x. \<not>accp R x \<longrightarrow> R (f x) x \<and> \<not>accp R (f x)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
31  | 
by (rule choice, auto simp:not_acc_smaller)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
33  | 
then obtain f where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
34  | 
in_R: "\<And>x. \<not>accp R x \<Longrightarrow> R (f x) x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
35  | 
and nia: "\<And>x. \<not>accp R x \<Longrightarrow> \<not>accp R (f x)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
36  | 
by blast  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
38  | 
let ?s = "\<lambda>i. (f ^ i) x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
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  | 
fix i  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
42  | 
have "\<not>accp R (?s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
43  | 
by (induct i) (auto simp:nia `\<not>accp R x`)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
44  | 
hence "R (f (?s i)) (?s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
45  | 
by (rule in_R)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
46  | 
}  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
48  | 
hence "idseq R ?s x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
49  | 
unfolding idseq_def  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
50  | 
by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
52  | 
thus ?thesis by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
53  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
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  | 
types ('a, 'q) cdesc =
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
60  | 
  "('q \<Rightarrow> bool) \<times> ('q \<Rightarrow> 'a) \<times>('q \<Rightarrow> 'a)"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
63  | 
fun in_cdesc :: "('a, 'q) cdesc \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
64  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
65  | 
"in_cdesc (\<Gamma>, r, l) x y = (\<exists>q. x = r q \<and> y = l q \<and> \<Gamma> q)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
66  | 
|
| 25765 | 67  | 
primrec mk_rel :: "('a, 'q) cdesc list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
68  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
69  | 
"mk_rel [] x y = False"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
70  | 
| "mk_rel (c#cs) x y =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
71  | 
(in_cdesc c x y \<or> mk_rel cs x y)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
74  | 
lemma some_rd:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
75  | 
assumes "mk_rel rds x y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
76  | 
shows "\<exists>rd\<in>set rds. in_cdesc rd x y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
77  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
78  | 
by (induct rds) (auto simp:in_cdesc_def)  | 
| 
 
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  | 
(* from a value sequence, get a sequence of rds *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
82  | 
lemma ex_cs:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
83  | 
assumes idseq: "idseq (mk_rel rds) s x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
84  | 
shows "\<exists>cs. \<forall>i. cs i \<in> set rds \<and> in_cdesc (cs i) (s (Suc i)) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
85  | 
proof -  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
86  | 
from idseq  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
87  | 
have a: "\<forall>i. \<exists>rd \<in> set rds. in_cdesc rd (s (Suc i)) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
88  | 
by (auto simp:idseq_def intro:some_rd)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
90  | 
show ?thesis  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
91  | 
by (rule choice) (insert a, blast)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
92  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
94  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
95  | 
types 'a measures = "nat \<Rightarrow> 'a \<Rightarrow> nat"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
97  | 
fun stepP :: "('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> 
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
98  | 
  ('a \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
99  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
100  | 
"stepP (\<Gamma>1,r1,l1) (\<Gamma>2,r2,l2) m1 m2 R  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
101  | 
= (\<forall>q\<^isub>1 q\<^isub>2. \<Gamma>1 q\<^isub>1 \<and> \<Gamma>2 q\<^isub>2 \<and> r1 q\<^isub>1 = l2 q\<^isub>2  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
102  | 
\<longrightarrow> R (m2 (l2 q\<^isub>2)) ((m1 (l1 q\<^isub>1))))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
105  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
106  | 
  decr :: "('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> 
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
107  | 
  ('a \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
108  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
109  | 
"decr c1 c2 m1 m2 = stepP c1 c2 m1 m2 (op <)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
111  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
112  | 
  decreq :: "('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> 
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
113  | 
  ('a \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
114  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
115  | 
"decreq c1 c2 m1 m2 = stepP c1 c2 m1 m2 (op \<le>)"  | 
| 
 
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  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
118  | 
  no_step :: "('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
119  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
120  | 
"no_step c1 c2 = stepP c1 c2 (\<lambda>x. 0) (\<lambda>x. 0) (\<lambda>x y. False)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
121  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
122  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
124  | 
lemma decr_in_cdesc:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
125  | 
assumes "in_cdesc RD1 y x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
126  | 
assumes "in_cdesc RD2 z y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
127  | 
assumes "decr RD1 RD2 m1 m2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
128  | 
shows "m2 y < m1 x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
129  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
130  | 
by (cases RD1, cases RD2, auto simp:decr_def)  | 
| 
 
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  | 
lemma decreq_in_cdesc:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
133  | 
assumes "in_cdesc RD1 y x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
134  | 
assumes "in_cdesc RD2 z y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
135  | 
assumes "decreq RD1 RD2 m1 m2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
136  | 
shows "m2 y \<le> m1 x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
137  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
138  | 
by (cases RD1, cases RD2, auto simp:decreq_def)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
141  | 
lemma no_inf_desc_nat_sequence:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
142  | 
fixes s :: "nat \<Rightarrow> nat"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
143  | 
assumes leq: "\<And>i. n \<le> i \<Longrightarrow> s (Suc i) \<le> s i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
144  | 
assumes less: "\<exists>\<^sub>\<infinity>i. s (Suc i) < s i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
145  | 
shows False  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
146  | 
proof -  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
147  | 
  {
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
148  | 
fix i j:: nat  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
149  | 
assume "n \<le> i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
150  | 
assume "i \<le> j"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
151  | 
	{
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
152  | 
fix k  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
153  | 
have "s (i + k) \<le> s i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
154  | 
proof (induct k)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
155  | 
case 0 thus ?case by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
156  | 
next  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
157  | 
case (Suc k)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
158  | 
with leq[of "i + k"] `n \<le> i`  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
159  | 
show ?case by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
160  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
161  | 
}  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
162  | 
from this[of "j - i"] `n \<le> i` `i \<le> j`  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
163  | 
have "s j \<le> s i" by auto  | 
| 
 
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  | 
note decr = this  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
167  | 
let ?min = "LEAST x. x \<in> range (\<lambda>i. s (n + i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
168  | 
have "?min \<in> range (\<lambda>i. s (n + i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
169  | 
by (rule LeastI) auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
170  | 
then obtain k where min: "?min = s (n + k)" by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
171  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
172  | 
from less  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
173  | 
obtain k' where "n + k < k'"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
174  | 
and "s (Suc k') < s k'"  | 
| 27417 | 175  | 
unfolding INFM_nat by auto  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
176  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
177  | 
with decr[of "n + k" k'] min  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
178  | 
have "s (Suc k') < ?min" by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
179  | 
moreover from `n + k < k'`  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
180  | 
have "s (Suc k') = s (n + (Suc k' - n))" by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
181  | 
ultimately  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
182  | 
show False using not_less_Least by blast  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
183  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
185  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
186  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
187  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
188  | 
  approx :: "nat scg \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc 
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
189  | 
\<Rightarrow> 'a measures \<Rightarrow> 'a measures \<Rightarrow> bool"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
190  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
191  | 
"approx G C C' M M'  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
192  | 
= (\<forall>i j. (dsc G i j \<longrightarrow> decr C C' (M i) (M' j))  | 
| 26513 | 193  | 
\<and>(eqp G i j \<longrightarrow> decreq C C' (M i) (M' j)))"  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
194  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
195  | 
|
| 
 
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  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
198  | 
(* Unfolding "approx" for finite graphs *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
199  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
200  | 
lemma approx_empty:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
201  | 
  "approx (Graph {}) c1 c2 ms1 ms2"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
202  | 
unfolding approx_def has_edge_def dest_graph.simps by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
203  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
204  | 
lemma approx_less:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
205  | 
assumes "stepP c1 c2 (ms1 i) (ms2 j) (op <)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
206  | 
assumes "approx (Graph Es) c1 c2 ms1 ms2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
207  | 
shows "approx (Graph (insert (i, \<down>, j) Es)) c1 c2 ms1 ms2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
208  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
209  | 
unfolding approx_def has_edge_def dest_graph.simps decr_def  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
210  | 
by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
212  | 
lemma approx_leq:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
213  | 
assumes "stepP c1 c2 (ms1 i) (ms2 j) (op \<le>)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
214  | 
assumes "approx (Graph Es) c1 c2 ms1 ms2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
215  | 
shows "approx (Graph (insert (i, \<Down>, j) Es)) c1 c2 ms1 ms2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
216  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
217  | 
unfolding approx_def has_edge_def dest_graph.simps decreq_def  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
218  | 
by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
221  | 
lemma "approx (Graph {(1, \<down>, 2),(2, \<Down>, 3)}) c1 c2 ms1 ms2"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
222  | 
apply (intro approx_less approx_leq approx_empty)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
223  | 
oops  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
226  | 
(*  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
227  | 
fun  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
228  | 
  no_step :: "('a, 'q) cdesc \<Rightarrow> ('a, 'q) cdesc \<Rightarrow> bool"
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
229  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
230  | 
"no_step (\<Gamma>1, r1, l1) (\<Gamma>2, r2, l2) =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
231  | 
(\<forall>q\<^isub>1 q\<^isub>2. \<Gamma>1 q\<^isub>1 \<and> \<Gamma>2 q\<^isub>2 \<and> r1 q\<^isub>1 = l2 q\<^isub>2 \<longrightarrow> False)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
232  | 
*)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
233  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
234  | 
lemma no_stepI:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
235  | 
"stepP c1 c2 m1 m2 (\<lambda>x y. False)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
236  | 
\<Longrightarrow> no_step c1 c2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
237  | 
by (cases c1, cases c2) (auto simp: no_step_def)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
239  | 
definition  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
240  | 
  sound_int :: "nat acg \<Rightarrow> ('a, 'q) cdesc list 
 | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
241  | 
\<Rightarrow> 'a measures list \<Rightarrow> bool"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
242  | 
where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
243  | 
"sound_int \<A> RDs M =  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
244  | 
(\<forall>n<length RDs. \<forall>m<length RDs.  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
245  | 
no_step (RDs ! n) (RDs ! m) \<or>  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
246  | 
(\<exists>G. (\<A> \<turnstile> n \<leadsto>\<^bsup>G\<^esup> m) \<and> approx G (RDs ! n) (RDs ! m) (M ! n) (M ! m)))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
249  | 
(* The following are uses by the tactics *)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
250  | 
lemma length_simps: "length [] = 0" "length (x#xs) = Suc (length xs)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
251  | 
by auto  | 
| 
 
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  | 
lemma all_less_zero: "\<forall>n<(0::nat). P n"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
254  | 
by simp  | 
| 
 
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  | 
lemma all_less_Suc:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
257  | 
assumes Pk: "P k"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
258  | 
assumes Pn: "\<forall>n<k. P n"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
259  | 
shows "\<forall>n<Suc k. P n"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
260  | 
proof (intro allI impI)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
261  | 
fix n assume "n < Suc k"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
262  | 
show "P n"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
263  | 
proof (cases "n < k")  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
264  | 
case True with Pn show ?thesis by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
265  | 
next  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
266  | 
case False with `n < Suc k` have "n = k" by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
267  | 
with Pk show ?thesis by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
268  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
269  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
270  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
272  | 
lemma step_witness:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
273  | 
assumes "in_cdesc RD1 y x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
274  | 
assumes "in_cdesc RD2 z y"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
275  | 
shows "\<not> no_step RD1 RD2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
276  | 
using assms  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
277  | 
by (cases RD1, cases RD2) (auto simp:no_step_def)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
278  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
279  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
280  | 
theorem SCT_on_relations:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
281  | 
assumes R: "R = mk_rel RDs"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
282  | 
assumes sound: "sound_int \<A> RDs M"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
283  | 
assumes "SCT \<A>"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
284  | 
shows "\<forall>x. accp R x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
285  | 
proof (rule, rule classical)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
286  | 
fix x  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
287  | 
assume "\<not> accp R x"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
288  | 
with non_acc_has_idseq  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
289  | 
have "\<exists>s. idseq R s x" .  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
290  | 
then obtain s where "idseq R s x" ..  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
291  | 
hence "\<exists>cs. \<forall>i. cs i \<in> set RDs \<and>  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
292  | 
in_cdesc (cs i) (s (Suc i)) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
293  | 
unfolding R by (rule ex_cs)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
294  | 
then obtain cs where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
295  | 
[simp]: "\<And>i. cs i \<in> set RDs"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
296  | 
and ird[simp]: "\<And>i. in_cdesc (cs i) (s (Suc i)) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
297  | 
by blast  | 
| 
 
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  | 
let ?cis = "\<lambda>i. index_of RDs (cs i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
300  | 
have "\<forall>i. \<exists>G. (\<A> \<turnstile> ?cis i \<leadsto>\<^bsup>G\<^esup> (?cis (Suc i)))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
301  | 
\<and> approx G (RDs ! ?cis i) (RDs ! ?cis (Suc i))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
302  | 
(M ! ?cis i) (M ! ?cis (Suc i))" (is "\<forall>i. \<exists>G. ?P i G")  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
303  | 
proof  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
304  | 
fix i  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
305  | 
let ?n = "?cis i" and ?n' = "?cis (Suc i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
307  | 
have "in_cdesc (RDs ! ?n) (s (Suc i)) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
308  | 
"in_cdesc (RDs ! ?n') (s (Suc (Suc i))) (s (Suc i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
309  | 
by (simp_all add:index_of_member)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
310  | 
with step_witness  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
311  | 
have "\<not> no_step (RDs ! ?n) (RDs ! ?n')" .  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
312  | 
moreover have  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
313  | 
"?n < length RDs"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
314  | 
"?n' < length RDs"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
315  | 
by (simp_all add:index_of_length[symmetric])  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
316  | 
ultimately  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
317  | 
obtain G  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
318  | 
where "\<A> \<turnstile> ?n \<leadsto>\<^bsup>G\<^esup> ?n'"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
319  | 
and "approx G (RDs ! ?n) (RDs ! ?n') (M ! ?n) (M ! ?n')"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
320  | 
using sound  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
321  | 
unfolding sound_int_def by auto  | 
| 
 
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  | 
thus "\<exists>G. ?P i G" by blast  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
324  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
325  | 
with choice  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
326  | 
have "\<exists>Gs. \<forall>i. ?P i (Gs i)" .  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
327  | 
then obtain Gs where  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
328  | 
A: "\<And>i. \<A> \<turnstile> ?cis i \<leadsto>\<^bsup>(Gs i)\<^esup> (?cis (Suc i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
329  | 
and B: "\<And>i. approx (Gs i) (RDs ! ?cis i) (RDs ! ?cis (Suc i))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
330  | 
(M ! ?cis i) (M ! ?cis (Suc i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
331  | 
by blast  | 
| 
 
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  | 
let ?p = "\<lambda>i. (?cis i, Gs i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
334  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
335  | 
from A have "has_ipath \<A> ?p"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
336  | 
unfolding has_ipath_def  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
337  | 
by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
338  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
339  | 
with `SCT \<A>` SCT_def  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
340  | 
obtain th where "is_desc_thread th ?p"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
341  | 
by auto  | 
| 
 
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  | 
then obtain n  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
344  | 
where fr: "\<forall>i\<ge>n. eqlat ?p th i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
345  | 
and inf: "\<exists>\<^sub>\<infinity>i. descat ?p th i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
346  | 
unfolding is_desc_thread_def by auto  | 
| 
 
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  | 
from B  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
349  | 
have approx:  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
350  | 
"\<And>i. approx (Gs i) (cs i) (cs (Suc i))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
351  | 
(M ! ?cis i) (M ! ?cis (Suc i))"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
352  | 
by (simp add:index_of_member)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
353  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
354  | 
let ?seq = "\<lambda>i. (M ! ?cis i) (th i) (s i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
355  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
356  | 
have "\<And>i. n < i \<Longrightarrow> ?seq (Suc i) \<le> ?seq i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
357  | 
proof -  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
358  | 
fix i  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
359  | 
let ?q1 = "th i" and ?q2 = "th (Suc i)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
360  | 
assume "n < i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
361  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
362  | 
with fr have "eqlat ?p th i" by simp  | 
| 26513 | 363  | 
hence "dsc (Gs i) ?q1 ?q2 \<or> eqp (Gs i) ?q1 ?q2"  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
364  | 
by simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
365  | 
thus "?seq (Suc i) \<le> ?seq i"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
366  | 
proof  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
367  | 
assume "dsc (Gs i) ?q1 ?q2"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
368  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
369  | 
with approx  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
370  | 
have a:"decr (cs i) (cs (Suc i))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
371  | 
((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
372  | 
unfolding approx_def by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
373  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
374  | 
show ?thesis  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
375  | 
apply (rule less_imp_le)  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
376  | 
apply (rule decr_in_cdesc[of _ "s (Suc i)" "s i"])  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
377  | 
by (rule ird a)+  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
378  | 
next  | 
| 26513 | 379  | 
assume "eqp (Gs i) ?q1 ?q2"  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
380  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
381  | 
with approx  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
382  | 
have a:"decreq (cs i) (cs (Suc i))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
383  | 
((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
384  | 
unfolding approx_def by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
385  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
386  | 
show ?thesis  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
387  | 
apply (rule decreq_in_cdesc[of _ "s (Suc i)" "s i"])  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
388  | 
by (rule ird a)+  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
389  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
390  | 
qed  | 
| 27417 | 391  | 
moreover have "\<exists>\<^sub>\<infinity>i. ?seq (Suc i) < ?seq i" unfolding INFM_nat  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
392  | 
proof  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
393  | 
fix i  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
394  | 
from inf obtain j where "i < j" and d: "descat ?p th j"  | 
| 27417 | 395  | 
unfolding INFM_nat by auto  | 
| 
25314
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
396  | 
let ?q1 = "th j" and ?q2 = "th (Suc j)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
397  | 
from d have "dsc (Gs j) ?q1 ?q2" by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
399  | 
with approx  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
400  | 
have a:"decr (cs j) (cs (Suc j))  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
401  | 
((M ! ?cis j) ?q1) ((M ! ?cis (Suc j)) ?q2)"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
402  | 
unfolding approx_def by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
403  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
404  | 
have "?seq (Suc j) < ?seq j"  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
405  | 
apply (rule decr_in_cdesc[of _ "s (Suc j)" "s j"])  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
406  | 
by (rule ird a)+  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
407  | 
with `i < j`  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
408  | 
show "\<exists>j. i < j \<and> ?seq (Suc j) < ?seq j" by auto  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
409  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
410  | 
ultimately have False  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
411  | 
by (rule no_inf_desc_nat_sequence[of "Suc n"]) simp  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
412  | 
thus "accp R x" ..  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
413  | 
qed  | 
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 
krauss 
parents:  
diff
changeset
 | 
415  | 
end  |