author | haftmann |
Wed, 10 Jun 2009 15:04:33 +0200 | |
changeset 31604 | eb2f9d709296 |
parent 28132 | 236e07d8821e |
child 31979 | 09f65e860bdb |
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_Theorem.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 "Proof of the Size-Change Principle" |
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 Correctness |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
9 |
imports Main Ramsey 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 |
subsection {* Auxiliary definitions *} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
13 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
14 |
definition is_thread :: "nat \<Rightarrow> 'a sequence \<Rightarrow> ('a, 'a scg) ipath \<Rightarrow> bool" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
15 |
where |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
16 |
"is_thread n \<theta> p = (\<forall>i\<ge>n. eqlat p \<theta> i)" |
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 |
definition is_fthread :: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
19 |
"'a sequence \<Rightarrow> ('a, 'a scg) ipath \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
20 |
where |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
21 |
"is_fthread \<theta> mp i j = (\<forall>k\<in>{i..<j}. eqlat mp \<theta> k)" |
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 |
definition is_desc_fthread :: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
24 |
"'a sequence \<Rightarrow> ('a, 'a scg) ipath \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
25 |
where |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
26 |
"is_desc_fthread \<theta> mp i j = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
27 |
(is_fthread \<theta> mp i j \<and> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
28 |
(\<exists>k\<in>{i..<j}. descat mp \<theta> k))" |
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 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
31 |
"has_fth p i j n m = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
32 |
(\<exists>\<theta>. is_fthread \<theta> p i j \<and> \<theta> i = n \<and> \<theta> j = m)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
33 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
34 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
35 |
"has_desc_fth p i j n m = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
36 |
(\<exists>\<theta>. is_desc_fthread \<theta> p i j \<and> \<theta> i = n \<and> \<theta> j = m)" |
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 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
39 |
subsection {* Everything is finite *} |
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 |
lemma finite_range: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
42 |
fixes f :: "nat \<Rightarrow> 'a" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
43 |
assumes fin: "finite (range f)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
44 |
shows "\<exists>x. \<exists>\<^sub>\<infinity>i. f i = x" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
45 |
proof (rule classical) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
46 |
assume "\<not>(\<exists>x. \<exists>\<^sub>\<infinity>i. f i = x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
47 |
hence "\<forall>x. \<exists>j. \<forall>i>j. f i \<noteq> x" |
27417 | 48 |
unfolding INFM_nat by blast |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
49 |
with choice |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
50 |
have "\<exists>j. \<forall>x. \<forall>i>(j x). f i \<noteq> x" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
51 |
then obtain j where |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
52 |
neq: "\<And>x i. j x < i \<Longrightarrow> f i \<noteq> x" by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
53 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
54 |
from fin have "finite (range (j o f))" |
27418 | 55 |
by (auto simp:comp_def range_composition) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
56 |
with finite_nat_bounded |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
57 |
obtain m where "range (j o f) \<subseteq> {..<m}" by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
58 |
hence "j (f m) < m" unfolding comp_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
59 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
60 |
with neq[of "f m" m] show ?thesis by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
61 |
qed |
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 |
lemma finite_range_ignore_prefix: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
64 |
fixes f :: "nat \<Rightarrow> 'a" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
65 |
assumes fA: "finite A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
66 |
assumes inA: "\<forall>x\<ge>n. f x \<in> A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
67 |
shows "finite (range f)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
68 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
69 |
have a: "UNIV = {0 ..< (n::nat)} \<union> { x. n \<le> x }" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
70 |
have b: "range f = f ` {0 ..< n} \<union> f ` { x. n \<le> x }" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
71 |
(is "\<dots> = ?A \<union> ?B") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
72 |
by (unfold a) (simp add:image_Un) |
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 |
have "finite ?A" by (rule finite_imageI) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
75 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
76 |
from inA have "?B \<subseteq> A" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
77 |
from this fA have "finite ?B" by (rule finite_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
78 |
ultimately show ?thesis using b by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
79 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
80 |
|
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 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
83 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
84 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
85 |
"finite_graph G = finite (dest_graph G)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
86 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
87 |
"all_finite G = (\<forall>n H m. has_edge G n H m \<longrightarrow> finite_graph H)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
88 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
89 |
"finite_acg A = (finite_graph A \<and> all_finite A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
90 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
91 |
"nodes G = fst ` dest_graph G \<union> snd ` snd ` dest_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
92 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
93 |
"edges G = fst ` snd ` dest_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
94 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
95 |
"smallnodes G = \<Union>(nodes ` edges G)" |
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 |
lemma thread_image_nodes: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
98 |
assumes th: "is_thread n \<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
99 |
shows "\<forall>i\<ge>n. \<theta> i \<in> nodes (snd (p i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
100 |
using prems |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
101 |
unfolding is_thread_def has_edge_def nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
102 |
by force |
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 |
lemma finite_nodes: "finite_graph G \<Longrightarrow> finite (nodes G)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
105 |
unfolding finite_graph_def nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
106 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
107 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
108 |
lemma nodes_subgraph: "A \<le> B \<Longrightarrow> nodes A \<subseteq> nodes B" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
109 |
unfolding graph_leq_def nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
110 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
111 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
112 |
lemma finite_edges: "finite_graph G \<Longrightarrow> finite (edges G)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
113 |
unfolding finite_graph_def edges_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
114 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
115 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
116 |
lemma edges_sum[simp]: "edges (A + B) = edges A \<union> edges B" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
117 |
unfolding edges_def graph_plus_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
118 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
119 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
120 |
lemma nodes_sum[simp]: "nodes (A + B) = nodes A \<union> nodes B" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
121 |
unfolding nodes_def graph_plus_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
122 |
by auto |
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 finite_acg_subset: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
125 |
"A \<le> B \<Longrightarrow> finite_acg B \<Longrightarrow> finite_acg A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
126 |
unfolding finite_acg_def finite_graph_def all_finite_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
127 |
has_edge_def graph_leq_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
128 |
by (auto elim:finite_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
129 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
130 |
lemma scg_finite: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
131 |
fixes G :: "'a scg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
132 |
assumes fin: "finite (nodes G)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
133 |
shows "finite_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
134 |
unfolding finite_graph_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
135 |
proof (rule finite_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
136 |
show "dest_graph G \<subseteq> nodes G \<times> UNIV \<times> nodes G" (is "_ \<subseteq> ?P") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
137 |
unfolding nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
138 |
by force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
139 |
show "finite ?P" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
140 |
by (intro finite_cartesian_product fin finite) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
141 |
qed |
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 |
lemma smallnodes_sum[simp]: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
144 |
"smallnodes (A + B) = smallnodes A \<union> smallnodes B" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
145 |
unfolding smallnodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
146 |
by auto |
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 |
lemma in_smallnodes: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
149 |
fixes A :: "'a acg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
150 |
assumes e: "has_edge A x G y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
151 |
shows "nodes G \<subseteq> smallnodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
152 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
153 |
have "fst (snd (x, G, y)) \<in> fst ` snd ` dest_graph A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
154 |
unfolding has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
155 |
by (rule imageI)+ (rule e[unfolded has_edge_def]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
156 |
then have "G \<in> edges A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
157 |
unfolding edges_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
158 |
thus ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
159 |
unfolding smallnodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
160 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
161 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
162 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
163 |
lemma finite_smallnodes: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
164 |
assumes fA: "finite_acg A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
165 |
shows "finite (smallnodes A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
166 |
unfolding smallnodes_def edges_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
167 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
168 |
from fA |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
169 |
show "finite (nodes ` fst ` snd ` dest_graph A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
170 |
unfolding finite_acg_def finite_graph_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
171 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
172 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
173 |
fix M assume "M \<in> nodes ` fst ` snd ` dest_graph A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
174 |
then obtain n G m |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
175 |
where M: "M = nodes G" and nGm: "(n,G,m) \<in> dest_graph A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
176 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
177 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
178 |
from fA |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
179 |
have "all_finite A" unfolding finite_acg_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
180 |
with nGm have "finite_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
181 |
unfolding all_finite_def has_edge_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
182 |
with finite_nodes |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
183 |
show "finite M" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
184 |
unfolding finite_graph_def M . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
185 |
qed |
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 |
lemma nodes_tcl: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
188 |
"nodes (tcl A) = nodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
189 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
190 |
show "nodes A \<subseteq> nodes (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
191 |
apply (rule nodes_subgraph) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
192 |
by (subst tcl_unfold_right) simp |
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 |
show "nodes (tcl A) \<subseteq> nodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
195 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
196 |
fix x assume "x \<in> nodes (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
197 |
then obtain z G y |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
198 |
where z: "z \<in> dest_graph (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
199 |
and dis: "z = (x, G, y) \<or> z = (y, G, x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
200 |
unfolding nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
201 |
by auto force+ |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
202 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
203 |
from dis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
204 |
show "x \<in> nodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
205 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
206 |
assume "z = (x, G, y)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
207 |
with z have "has_edge (tcl A) x G y" unfolding has_edge_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
208 |
then obtain n where "n > 0 " and An: "has_edge (A ^ n) x G y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
209 |
unfolding in_tcl by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
210 |
then obtain n' where "n = Suc n'" by (cases n, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
211 |
hence "A ^ n = A * A ^ n'" by (simp add:power_Suc) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
212 |
with An obtain e k |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
213 |
where "has_edge A x e k" by (auto simp:in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
214 |
thus "x \<in> nodes A" unfolding has_edge_def nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
215 |
by force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
216 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
217 |
assume "z = (y, G, x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
218 |
with z have "has_edge (tcl A) y G x" unfolding has_edge_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
219 |
then obtain n where "n > 0 " and An: "has_edge (A ^ n) y G x" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
220 |
unfolding in_tcl by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
221 |
then obtain n' where "n = Suc n'" by (cases n, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
222 |
hence "A ^ n = A ^ n' * A" by (simp add:power_Suc power_commutes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
223 |
with An obtain e k |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
224 |
where "has_edge A k e x" by (auto simp:in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
225 |
thus "x \<in> nodes A" unfolding has_edge_def nodes_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
226 |
by force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
227 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
228 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
229 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
230 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
231 |
lemma smallnodes_tcl: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
232 |
fixes A :: "'a acg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
233 |
shows "smallnodes (tcl A) = smallnodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
234 |
proof (intro equalityI subsetI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
235 |
fix n assume "n \<in> smallnodes (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
236 |
then obtain x G y where edge: "has_edge (tcl A) x G y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
237 |
and "n \<in> nodes G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
238 |
unfolding smallnodes_def edges_def has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
239 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
240 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
241 |
from `n \<in> nodes G` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
242 |
have "n \<in> fst ` dest_graph G \<or> n \<in> snd ` snd ` dest_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
243 |
(is "?A \<or> ?B") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
244 |
unfolding nodes_def by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
245 |
thus "n \<in> smallnodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
246 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
247 |
assume ?A |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
248 |
then obtain m e where A: "has_edge G n e m" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
249 |
unfolding has_edge_def by auto |
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 |
have "tcl A = A * star A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
252 |
unfolding tcl_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
253 |
by (simp add: star_commute[of A A A, simplified]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
254 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
255 |
with edge |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
256 |
have "has_edge (A * star A) x G y" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
257 |
then obtain H H' z |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
258 |
where AH: "has_edge A x H z" and G: "G = H * H'" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
259 |
by (auto simp:in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
260 |
from A |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
261 |
obtain m' e' where "has_edge H n e' m'" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
262 |
by (auto simp:G in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
263 |
hence "n \<in> nodes H" unfolding nodes_def has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
264 |
by force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
265 |
with in_smallnodes[OF AH] show "n \<in> smallnodes A" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
266 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
267 |
assume ?B |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
268 |
then obtain m e where B: "has_edge G m e n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
269 |
unfolding has_edge_def by auto |
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 |
with edge |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
272 |
have "has_edge (star A * A) x G y" by (simp add:tcl_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
273 |
then obtain H H' z |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
274 |
where AH': "has_edge A z H' y" and G: "G = H * H'" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
275 |
by (auto simp:in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
276 |
from B |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
277 |
obtain m' e' where "has_edge H' m' e' n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
278 |
by (auto simp:G in_grcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
279 |
hence "n \<in> nodes H'" unfolding nodes_def has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
280 |
by force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
281 |
with in_smallnodes[OF AH'] show "n \<in> smallnodes A" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
282 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
283 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
284 |
fix x assume "x \<in> smallnodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
285 |
then show "x \<in> smallnodes (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
286 |
by (subst tcl_unfold_right) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
287 |
qed |
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 |
lemma finite_nodegraphs: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
290 |
assumes F: "finite F" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
291 |
shows "finite { G::'a scg. nodes G \<subseteq> F }" (is "finite ?P") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
292 |
proof (rule finite_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
293 |
show "?P \<subseteq> Graph ` (Pow (F \<times> UNIV \<times> F))" (is "?P \<subseteq> ?Q") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
294 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
295 |
fix x assume xP: "x \<in> ?P" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
296 |
obtain S where x[simp]: "x = Graph S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
297 |
by (cases x) auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
298 |
from xP |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
299 |
show "x \<in> ?Q" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
300 |
apply (simp add:nodes_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
301 |
apply (rule imageI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
302 |
apply (rule PowI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
303 |
apply force |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
304 |
done |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
305 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
306 |
show "finite ?Q" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
307 |
by (auto intro:finite_imageI finite_cartesian_product F finite) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
308 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
309 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
310 |
lemma finite_graphI: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
311 |
fixes A :: "'a acg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
312 |
assumes fin: "finite (nodes A)" "finite (smallnodes A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
313 |
shows "finite_graph A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
314 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
315 |
obtain S where A[simp]: "A = Graph S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
316 |
by (cases A) auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
317 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
318 |
have "finite S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
319 |
proof (rule finite_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
320 |
show "S \<subseteq> nodes A \<times> { G::'a scg. nodes G \<subseteq> smallnodes A } \<times> nodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
321 |
(is "S \<subseteq> ?T") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
322 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
323 |
fix x assume xS: "x \<in> S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
324 |
obtain a b c where x[simp]: "x = (a, b, c)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
325 |
by (cases x) auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
326 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
327 |
then have edg: "has_edge A a b c" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
328 |
unfolding has_edge_def using xS |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
329 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
330 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
331 |
hence "a \<in> nodes A" "c \<in> nodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
332 |
unfolding nodes_def has_edge_def by force+ |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
333 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
334 |
from edg have "nodes b \<subseteq> smallnodes A" by (rule in_smallnodes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
335 |
hence "b \<in> { G :: 'a scg. nodes G \<subseteq> smallnodes A }" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
336 |
ultimately show "x \<in> ?T" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
337 |
qed |
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 |
show "finite ?T" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
340 |
by (intro finite_cartesian_product fin finite_nodegraphs) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
341 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
342 |
thus ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
343 |
unfolding finite_graph_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
344 |
qed |
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 |
lemma smallnodes_allfinite: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
348 |
fixes A :: "'a acg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
349 |
assumes fin: "finite (smallnodes A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
350 |
shows "all_finite A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
351 |
unfolding all_finite_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
352 |
proof (intro allI impI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
353 |
fix n H m assume "has_edge A n H m" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
354 |
then have "nodes H \<subseteq> smallnodes A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
355 |
by (rule in_smallnodes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
356 |
then have "finite (nodes H)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
357 |
by (rule finite_subset) (rule fin) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
358 |
thus "finite_graph H" by (rule scg_finite) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
359 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
360 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
361 |
lemma finite_tcl: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
362 |
fixes A :: "'a acg" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
363 |
shows "finite_acg (tcl A) \<longleftrightarrow> finite_acg A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
364 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
365 |
assume f: "finite_acg A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
366 |
from f have g: "finite_graph A" and "all_finite A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
367 |
unfolding finite_acg_def by auto |
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 |
from g have "finite (nodes A)" by (rule finite_nodes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
370 |
then have "finite (nodes (tcl A))" unfolding nodes_tcl . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
371 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
372 |
from f have "finite (smallnodes A)" by (rule finite_smallnodes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
373 |
then have fs: "finite (smallnodes (tcl A))" unfolding smallnodes_tcl . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
374 |
ultimately |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
375 |
have "finite_graph (tcl A)" by (rule finite_graphI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
376 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
377 |
moreover from fs have "all_finite (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
378 |
by (rule smallnodes_allfinite) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
379 |
ultimately show "finite_acg (tcl A)" unfolding finite_acg_def .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
380 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
381 |
assume a: "finite_acg (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
382 |
have "A \<le> tcl A" by (rule less_tcl) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
383 |
thus "finite_acg A" using a |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
384 |
by (rule finite_acg_subset) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
385 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
386 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
387 |
lemma finite_acg_empty: "finite_acg (Graph {})" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
388 |
unfolding finite_acg_def finite_graph_def all_finite_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
389 |
has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
390 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
391 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
392 |
lemma finite_acg_ins: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
393 |
assumes fA: "finite_acg (Graph A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
394 |
assumes fG: "finite G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
395 |
shows "finite_acg (Graph (insert (a, Graph G, b) A))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
396 |
using fA fG |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
397 |
unfolding finite_acg_def finite_graph_def all_finite_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
398 |
has_edge_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
399 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
400 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
401 |
lemmas finite_acg_simps = finite_acg_empty finite_acg_ins finite_graph_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
402 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
403 |
subsection {* Contraction and more *} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
404 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
405 |
abbreviation |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
406 |
"pdesc P == (fst P, prod P, end_node P)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
407 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
408 |
lemma pdesc_acgplus: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
409 |
assumes "has_ipath \<A> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
410 |
and "i < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
411 |
shows "has_edge (tcl \<A>) (fst (p\<langle>i,j\<rangle>)) (prod (p\<langle>i,j\<rangle>)) (end_node (p\<langle>i,j\<rangle>))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
412 |
unfolding plus_paths |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
413 |
apply (rule exI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
414 |
apply (insert prems) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
415 |
by (auto intro:sub_path_is_path[of "\<A>" p i j] simp:sub_path_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
416 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
417 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
418 |
lemma combine_fthreads: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
419 |
assumes range: "i < j" "j \<le> k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
420 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
421 |
"has_fth p i k m r = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
422 |
(\<exists>n. has_fth p i j m n \<and> has_fth p j k n r)" (is "?L = ?R") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
423 |
proof (intro iffI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
424 |
assume "?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
425 |
then obtain \<theta> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
426 |
where "is_fthread \<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
427 |
and [simp]: "\<theta> i = m" "\<theta> k = r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
428 |
by (auto simp:has_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
429 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
430 |
with range |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
431 |
have "is_fthread \<theta> p i j" and "is_fthread \<theta> p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
432 |
by (auto simp:is_fthread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
433 |
hence "has_fth p i j m (\<theta> j)" and "has_fth p j k (\<theta> j) r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
434 |
by (auto simp:has_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
435 |
thus "?R" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
436 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
437 |
assume "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
438 |
then obtain n \<theta>1 \<theta>2 |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
439 |
where ths: "is_fthread \<theta>1 p i j" "is_fthread \<theta>2 p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
440 |
and [simp]: "\<theta>1 i = m" "\<theta>1 j = n" "\<theta>2 j = n" "\<theta>2 k = r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
441 |
by (auto simp:has_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
442 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
443 |
let ?\<theta> = "(\<lambda>i. if i < j then \<theta>1 i else \<theta>2 i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
444 |
have "is_fthread ?\<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
445 |
unfolding is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
446 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
447 |
fix l assume range: "l \<in> {i..<k}" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
448 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
449 |
show "eqlat p ?\<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
450 |
proof (cases rule:three_cases) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
451 |
assume "Suc l < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
452 |
with ths range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
453 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
454 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
455 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
456 |
assume "Suc l = j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
457 |
hence "l < j" "\<theta>2 (Suc l) = \<theta>1 (Suc l)" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
458 |
with ths range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
459 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
460 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
461 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
462 |
assume "j \<le> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
463 |
with ths range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
464 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
465 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
466 |
qed arith |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
467 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
468 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
469 |
have "?\<theta> i = m" "?\<theta> k = r" using range by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
470 |
ultimately show "has_fth p i k m r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
471 |
by (auto simp:has_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
472 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
473 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
474 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
475 |
lemma desc_is_fthread: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
476 |
"is_desc_fthread \<theta> p i k \<Longrightarrow> is_fthread \<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
477 |
unfolding is_desc_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
478 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
479 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
480 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
481 |
lemma combine_dfthreads: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
482 |
assumes range: "i < j" "j \<le> k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
483 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
484 |
"has_desc_fth p i k m r = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
485 |
(\<exists>n. (has_desc_fth p i j m n \<and> has_fth p j k n r) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
486 |
\<or> (has_fth p i j m n \<and> has_desc_fth p j k n r))" (is "?L = ?R") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
487 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
488 |
assume "?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
489 |
then obtain \<theta> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
490 |
where desc: "is_desc_fthread \<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
491 |
and [simp]: "\<theta> i = m" "\<theta> k = r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
492 |
by (auto simp:has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
493 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
494 |
hence "is_fthread \<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
495 |
by (simp add: desc_is_fthread) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
496 |
with range have fths: "is_fthread \<theta> p i j" "is_fthread \<theta> p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
497 |
unfolding is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
498 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
499 |
hence hfths: "has_fth p i j m (\<theta> j)" "has_fth p j k (\<theta> j) r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
500 |
by (auto simp:has_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
501 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
502 |
from desc obtain l |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
503 |
where "i \<le> l" "l < k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
504 |
and "descat p \<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
505 |
by (auto simp:is_desc_fthread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
506 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
507 |
with fths |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
508 |
have "is_desc_fthread \<theta> p i j \<or> is_desc_fthread \<theta> p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
509 |
unfolding is_desc_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
510 |
by (cases "l < j") auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
511 |
hence "has_desc_fth p i j m (\<theta> j) \<or> has_desc_fth p j k (\<theta> j) r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
512 |
by (auto simp:has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
513 |
with hfths show ?R |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
514 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
515 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
516 |
assume "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
517 |
then obtain n \<theta>1 \<theta>2 |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
518 |
where "(is_desc_fthread \<theta>1 p i j \<and> is_fthread \<theta>2 p j k) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
519 |
\<or> (is_fthread \<theta>1 p i j \<and> is_desc_fthread \<theta>2 p j k)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
520 |
and [simp]: "\<theta>1 i = m" "\<theta>1 j = n" "\<theta>2 j = n" "\<theta>2 k = r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
521 |
by (auto simp:has_fth_def has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
522 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
523 |
hence ths2: "is_fthread \<theta>1 p i j" "is_fthread \<theta>2 p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
524 |
and dths: "is_desc_fthread \<theta>1 p i j \<or> is_desc_fthread \<theta>2 p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
525 |
by (auto simp:desc_is_fthread) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
526 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
527 |
let ?\<theta> = "(\<lambda>i. if i < j then \<theta>1 i else \<theta>2 i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
528 |
have "is_fthread ?\<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
529 |
unfolding is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
530 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
531 |
fix l assume range: "l \<in> {i..<k}" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
532 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
533 |
show "eqlat p ?\<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
534 |
proof (cases rule:three_cases) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
535 |
assume "Suc l < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
536 |
with ths2 range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
537 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
538 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
539 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
540 |
assume "Suc l = j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
541 |
hence "l < j" "\<theta>2 (Suc l) = \<theta>1 (Suc l)" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
542 |
with ths2 range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
543 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
544 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
545 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
546 |
assume "j \<le> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
547 |
with ths2 range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
548 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
549 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
550 |
qed arith |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
551 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
552 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
553 |
from dths |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
554 |
have "\<exists>l. i \<le> l \<and> l < k \<and> descat p ?\<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
555 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
556 |
assume "is_desc_fthread \<theta>1 p i j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
557 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
558 |
then obtain l where range: "i \<le> l" "l < j" and "descat p \<theta>1 l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
559 |
unfolding is_desc_fthread_def Bex_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
560 |
hence "descat p ?\<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
561 |
by (cases "Suc l = j", auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
562 |
with `j \<le> k` and range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
563 |
by (rule_tac x="l" in exI, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
564 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
565 |
assume "is_desc_fthread \<theta>2 p j k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
566 |
then obtain l where range: "j \<le> l" "l < k" and "descat p \<theta>2 l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
567 |
unfolding is_desc_fthread_def Bex_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
568 |
with `i < j` have "descat p ?\<theta> l" "i \<le> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
569 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
570 |
with range show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
571 |
by (rule_tac x="l" in exI, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
572 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
573 |
ultimately have "is_desc_fthread ?\<theta> p i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
574 |
by (simp add: is_desc_fthread_def Bex_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
575 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
576 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
577 |
have "?\<theta> i = m" "?\<theta> k = r" using range by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
578 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
579 |
ultimately show "has_desc_fth p i k m r" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
580 |
by (auto simp:has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
581 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
582 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
583 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
584 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
585 |
lemma fth_single: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
586 |
"has_fth p i (Suc i) m n = eql (snd (p i)) m n" (is "?L = ?R") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
587 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
588 |
assume "?L" thus "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
589 |
unfolding is_fthread_def Ball_def has_fth_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
590 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
591 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
592 |
let ?\<theta> = "\<lambda>k. if k = i then m else n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
593 |
assume edge: "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
594 |
hence "is_fthread ?\<theta> p i (Suc i) \<and> ?\<theta> i = m \<and> ?\<theta> (Suc i) = n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
595 |
unfolding is_fthread_def Ball_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
596 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
597 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
598 |
thus "?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
599 |
unfolding has_fth_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
600 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
601 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
602 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
603 |
lemma desc_fth_single: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
604 |
"has_desc_fth p i (Suc i) m n = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
605 |
dsc (snd (p i)) m n" (is "?L = ?R") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
606 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
607 |
assume "?L" thus "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
608 |
unfolding is_desc_fthread_def has_desc_fth_def is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
609 |
Bex_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
610 |
by (elim exE conjE) (case_tac "k = i", auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
611 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
612 |
let ?\<theta> = "\<lambda>k. if k = i then m else n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
613 |
assume edge: "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
614 |
hence "is_desc_fthread ?\<theta> p i (Suc i) \<and> ?\<theta> i = m \<and> ?\<theta> (Suc i) = n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
615 |
unfolding is_desc_fthread_def is_fthread_def Ball_def Bex_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
616 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
617 |
thus "?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
618 |
unfolding has_desc_fth_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
619 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
620 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
621 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
622 |
lemma mk_eql: "(G \<turnstile> m \<leadsto>\<^bsup>e\<^esup> n) \<Longrightarrow> eql G m n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
623 |
by (cases e, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
624 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
625 |
lemma eql_scgcomp: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
626 |
"eql (G * H) m r = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
627 |
(\<exists>n. eql G m n \<and> eql H n r)" (is "?L = ?R") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
628 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
629 |
show "?L \<Longrightarrow> ?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
630 |
by (auto simp:in_grcomp intro!:mk_eql) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
631 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
632 |
assume "?R" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
633 |
then obtain n where l: "eql G m n" and r:"eql H n r" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
634 |
thus ?L |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
635 |
by (cases "dsc G m n") (auto simp:in_grcomp mult_sedge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
636 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
637 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
638 |
lemma desc_scgcomp: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
639 |
"dsc (G * H) m r = |
26513 | 640 |
(\<exists>n. (dsc G m n \<and> eql H n r) \<or> (eqp G m n \<and> dsc H n r))" (is "?L = ?R") |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
641 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
642 |
show "?R \<Longrightarrow> ?L" by (auto simp:in_grcomp mult_sedge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
643 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
644 |
assume "?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
645 |
thus ?R |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
646 |
by (auto simp:in_grcomp mult_sedge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
647 |
(case_tac "e", auto, case_tac "e'", auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
648 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
649 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
650 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
651 |
lemma has_fth_unfold: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
652 |
assumes "i < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
653 |
shows "has_fth p i j m n = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
654 |
(\<exists>r. has_fth p i (Suc i) m r \<and> has_fth p (Suc i) j r n)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
655 |
by (rule combine_fthreads) (insert `i < j`, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
656 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
657 |
lemma has_dfth_unfold: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
658 |
assumes range: "i < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
659 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
660 |
"has_desc_fth p i j m r = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
661 |
(\<exists>n. (has_desc_fth p i (Suc i) m n \<and> has_fth p (Suc i) j n r) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
662 |
\<or> (has_fth p i (Suc i) m n \<and> has_desc_fth p (Suc i) j n r))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
663 |
by (rule combine_dfthreads) (insert `i < j`, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
664 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
665 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
666 |
lemma Lemma7a: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
667 |
"i \<le> j \<Longrightarrow> has_fth p i j m n = eql (prod (p\<langle>i,j\<rangle>)) m n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
668 |
proof (induct i arbitrary: m rule:inc_induct) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
669 |
case base show ?case |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
670 |
unfolding has_fth_def is_fthread_def sub_path_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
671 |
by (auto simp:in_grunit one_sedge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
672 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
673 |
case (step i) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
674 |
note IH = `\<And>m. has_fth p (Suc i) j m n = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
675 |
eql (prod (p\<langle>Suc i,j\<rangle>)) m n` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
676 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
677 |
have "has_fth p i j m n |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
678 |
= (\<exists>r. has_fth p i (Suc i) m r \<and> has_fth p (Suc i) j r n)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
679 |
by (rule has_fth_unfold[OF `i < j`]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
680 |
also have "\<dots> = (\<exists>r. has_fth p i (Suc i) m r |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
681 |
\<and> eql (prod (p\<langle>Suc i,j\<rangle>)) r n)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
682 |
by (simp only:IH) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
683 |
also have "\<dots> = (\<exists>r. eql (snd (p i)) m r |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
684 |
\<and> eql (prod (p\<langle>Suc i,j\<rangle>)) r n)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
685 |
by (simp only:fth_single) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
686 |
also have "\<dots> = eql (snd (p i) * prod (p\<langle>Suc i,j\<rangle>)) m n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
687 |
by (simp only:eql_scgcomp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
688 |
also have "\<dots> = eql (prod (p\<langle>i,j\<rangle>)) m n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
689 |
by (simp only:prod_unfold[OF `i < j`]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
690 |
finally show ?case . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
691 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
692 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
693 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
694 |
lemma Lemma7b: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
695 |
assumes "i \<le> j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
696 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
697 |
"has_desc_fth p i j m n = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
698 |
dsc (prod (p\<langle>i,j\<rangle>)) m n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
699 |
using prems |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
700 |
proof (induct i arbitrary: m rule:inc_induct) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
701 |
case base show ?case |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
702 |
unfolding has_desc_fth_def is_desc_fthread_def sub_path_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
703 |
by (auto simp:in_grunit one_sedge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
704 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
705 |
case (step i) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
706 |
thus ?case |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
707 |
by (simp only:prod_unfold desc_scgcomp desc_fth_single |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
708 |
has_dfth_unfold fth_single Lemma7a) auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
709 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
710 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
711 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
712 |
lemma descat_contract: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
713 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
714 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
715 |
"descat (contract s p) \<theta> i = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
716 |
has_desc_fth p (s i) (s (Suc i)) (\<theta> i) (\<theta> (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
717 |
by (simp add:Lemma7b increasing_weak contract_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
718 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
719 |
lemma eqlat_contract: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
720 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
721 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
722 |
"eqlat (contract s p) \<theta> i = |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
723 |
has_fth p (s i) (s (Suc i)) (\<theta> i) (\<theta> (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
724 |
by (auto simp:Lemma7a increasing_weak contract_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
725 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
726 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
727 |
subsubsection {* Connecting threads *} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
728 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
729 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
730 |
"connect s \<theta>s = (\<lambda>k. \<theta>s (section_of s k) k)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
731 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
732 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
733 |
lemma next_in_range: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
734 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
735 |
assumes a: "k \<in> section s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
736 |
shows "(Suc k \<in> section s i) \<or> (Suc k \<in> section s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
737 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
738 |
from a have "k < s (Suc i)" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
739 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
740 |
hence "Suc k < s (Suc i) \<or> Suc k = s (Suc i)" by arith |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
741 |
thus ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
742 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
743 |
assume "Suc k < s (Suc i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
744 |
with a have "Suc k \<in> section s i" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
745 |
thus ?thesis .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
746 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
747 |
assume eq: "Suc k = s (Suc i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
748 |
with increasing_strict have "Suc k < s (Suc (Suc i))" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
749 |
with eq have "Suc k \<in> section s (Suc i)" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
750 |
thus ?thesis .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
751 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
752 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
753 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
754 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
755 |
lemma connect_threads: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
756 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
757 |
assumes connected: "\<theta>s i (s (Suc i)) = \<theta>s (Suc i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
758 |
assumes fth: "is_fthread (\<theta>s i) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
759 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
760 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
761 |
"is_fthread (connect s \<theta>s) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
762 |
unfolding is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
763 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
764 |
fix k assume krng: "k \<in> section s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
765 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
766 |
with fth have eqlat: "eqlat p (\<theta>s i) k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
767 |
unfolding is_fthread_def by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
768 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
769 |
from krng and next_in_range |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
770 |
have "(Suc k \<in> section s i) \<or> (Suc k \<in> section s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
771 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
772 |
thus "eqlat p (connect s \<theta>s) k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
773 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
774 |
assume "Suc k \<in> section s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
775 |
with krng eqlat show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
776 |
unfolding connect_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
777 |
by (simp only:section_of_known `increasing s`) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
778 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
779 |
assume skrng: "Suc k \<in> section s (Suc i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
780 |
with krng have "Suc k = s (Suc i)" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
781 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
782 |
with krng skrng eqlat show ?thesis |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
783 |
unfolding connect_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
784 |
by (simp only:section_of_known connected[symmetric] `increasing s`) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
785 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
786 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
787 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
788 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
789 |
lemma connect_dthreads: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
790 |
assumes inc[simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
791 |
assumes connected: "\<theta>s i (s (Suc i)) = \<theta>s (Suc i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
792 |
assumes fth: "is_desc_fthread (\<theta>s i) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
793 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
794 |
shows |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
795 |
"is_desc_fthread (connect s \<theta>s) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
796 |
unfolding is_desc_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
797 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
798 |
show "is_fthread (connect s \<theta>s) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
799 |
apply (rule connect_threads) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
800 |
apply (insert fth) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
801 |
by (auto simp:connected is_desc_fthread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
802 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
803 |
from fth |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
804 |
obtain k where dsc: "descat p (\<theta>s i) k" and krng: "k \<in> section s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
805 |
unfolding is_desc_fthread_def by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
806 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
807 |
from krng and next_in_range |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
808 |
have "(Suc k \<in> section s i) \<or> (Suc k \<in> section s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
809 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
810 |
hence "descat p (connect s \<theta>s) k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
811 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
812 |
assume "Suc k \<in> section s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
813 |
with krng dsc show ?thesis unfolding connect_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
814 |
by (simp only:section_of_known inc) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
815 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
816 |
assume skrng: "Suc k \<in> section s (Suc i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
817 |
with krng have "Suc k = s (Suc i)" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
818 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
819 |
with krng skrng dsc show ?thesis unfolding connect_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
820 |
by (simp only:section_of_known connected[symmetric] inc) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
821 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
822 |
with krng show "\<exists>k\<in>section s i. descat p (connect s \<theta>s) k" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
823 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
824 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
825 |
lemma mk_inf_thread: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
826 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
827 |
assumes fths: "\<And>i. i > n \<Longrightarrow> is_fthread \<theta> p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
828 |
shows "is_thread (s (Suc n)) \<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
829 |
unfolding is_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
830 |
proof (intro allI impI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
831 |
fix j assume st: "s (Suc n) \<le> j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
832 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
833 |
let ?k = "section_of s j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
834 |
from in_section_of st |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
835 |
have rs: "j \<in> section s ?k" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
836 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
837 |
with st have "s (Suc n) < s (Suc ?k)" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
838 |
with increasing_bij have "n < ?k" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
839 |
with rs and fths[of ?k] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
840 |
show "eqlat p \<theta> j" by (simp add:is_fthread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
841 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
842 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
843 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
844 |
lemma mk_inf_desc_thread: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
845 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
846 |
assumes fths: "\<And>i. i > n \<Longrightarrow> is_fthread \<theta> p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
847 |
assumes fdths: "\<exists>\<^sub>\<infinity>i. is_desc_fthread \<theta> p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
848 |
shows "is_desc_thread \<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
849 |
unfolding is_desc_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
850 |
proof (intro exI conjI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
851 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
852 |
from mk_inf_thread[of s n \<theta> p] fths |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
853 |
show "\<forall>i\<ge>s (Suc n). eqlat p \<theta> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
854 |
by (fold is_thread_def) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
855 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
856 |
show "\<exists>\<^sub>\<infinity>l. descat p \<theta> l" |
27417 | 857 |
unfolding INFM_nat |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
858 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
859 |
fix i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
860 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
861 |
let ?k = "section_of s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
862 |
from fdths obtain j |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
863 |
where "?k < j" "is_desc_fthread \<theta> p (s j) (s (Suc j))" |
27417 | 864 |
unfolding INFM_nat by auto |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
865 |
then obtain l where "s j \<le> l" and desc: "descat p \<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
866 |
unfolding is_desc_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
867 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
868 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
869 |
have "i < s (Suc ?k)" by (rule section_of2) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
870 |
also have "\<dots> \<le> s j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
871 |
by (rule increasing_weak [OF `increasing s`]) (insert `?k < j`, arith) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
872 |
also note `\<dots> \<le> l` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
873 |
finally have "i < l" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
874 |
with desc |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
875 |
show "\<exists>l. i < l \<and> descat p \<theta> l" by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
876 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
877 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
878 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
879 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
880 |
lemma desc_ex_choice: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
881 |
assumes A: "((\<exists>n.\<forall>i\<ge>n. \<exists>x. P x i) \<and> (\<exists>\<^sub>\<infinity>i. \<exists>x. Q x i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
882 |
and imp: "\<And>x i. Q x i \<Longrightarrow> P x i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
883 |
shows "\<exists>xs. ((\<exists>n.\<forall>i\<ge>n. P (xs i) i) \<and> (\<exists>\<^sub>\<infinity>i. Q (xs i) i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
884 |
(is "\<exists>xs. ?Ps xs \<and> ?Qs xs") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
885 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
886 |
let ?w = "\<lambda>i. (if (\<exists>x. Q x i) then (SOME x. Q x i) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
887 |
else (SOME x. P x i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
888 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
889 |
from A |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
890 |
obtain n where P: "\<And>i. n \<le> i \<Longrightarrow> \<exists>x. P x i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
891 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
892 |
{ |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
893 |
fix i::'a assume "n \<le> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
894 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
895 |
have "P (?w i) i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
896 |
proof (cases "\<exists>x. Q x i") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
897 |
case True |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
898 |
hence "Q (?w i) i" by (auto intro:someI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
899 |
with imp show "P (?w i) i" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
900 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
901 |
case False |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
902 |
with P[OF `n \<le> i`] show "P (?w i) i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
903 |
by (auto intro:someI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
904 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
905 |
} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
906 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
907 |
hence "?Ps ?w" by (rule_tac x=n in exI) auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
908 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
909 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
910 |
from A have "\<exists>\<^sub>\<infinity>i. (\<exists>x. Q x i)" .. |
27417 | 911 |
hence "?Qs ?w" by (rule INFM_mono) (auto intro:someI) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
912 |
ultimately |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
913 |
show "?Ps ?w \<and> ?Qs ?w" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
914 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
915 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
916 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
917 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
918 |
lemma dthreads_join: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
919 |
assumes [simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
920 |
assumes dthread: "is_desc_thread \<theta> (contract s p)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
921 |
shows "\<exists>\<theta>s. desc (\<lambda>i. is_fthread (\<theta>s i) p (s i) (s (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
922 |
\<and> \<theta>s i (s i) = \<theta> i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
923 |
\<and> \<theta>s i (s (Suc i)) = \<theta> (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
924 |
(\<lambda>i. is_desc_fthread (\<theta>s i) p (s i) (s (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
925 |
\<and> \<theta>s i (s i) = \<theta> i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
926 |
\<and> \<theta>s i (s (Suc i)) = \<theta> (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
927 |
apply (rule desc_ex_choice) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
928 |
apply (insert dthread) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
929 |
apply (simp only:is_desc_thread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
930 |
apply (simp add:eqlat_contract) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
931 |
apply (simp add:descat_contract) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
932 |
apply (simp only:has_fth_def has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
933 |
by (auto simp:is_desc_fthread_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
934 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
935 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
936 |
|
28132 | 937 |
lemma INFM_drop_prefix: |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
938 |
"(\<exists>\<^sub>\<infinity>i::nat. i > n \<and> P i) = (\<exists>\<^sub>\<infinity>i. P i)" |
27417 | 939 |
apply (auto simp:INFM_nat) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
940 |
apply (drule_tac x = "max m n" in spec) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
941 |
apply (elim exE conjE) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
942 |
apply (rule_tac x = "na" in exI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
943 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
944 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
945 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
946 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
947 |
lemma contract_keeps_threads: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
948 |
assumes inc[simp]: "increasing s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
949 |
shows "(\<exists>\<theta>. is_desc_thread \<theta> p) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
950 |
\<longleftrightarrow> (\<exists>\<theta>. is_desc_thread \<theta> (contract s p))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
951 |
(is "?A \<longleftrightarrow> ?B") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
952 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
953 |
assume "?A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
954 |
then obtain \<theta> n |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
955 |
where fr: "\<forall>i\<ge>n. eqlat p \<theta> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
956 |
and ds: "\<exists>\<^sub>\<infinity>i. descat p \<theta> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
957 |
unfolding is_desc_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
958 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
959 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
960 |
let ?c\<theta> = "\<lambda>i. \<theta> (s i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
961 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
962 |
have "is_desc_thread ?c\<theta> (contract s p)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
963 |
unfolding is_desc_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
964 |
proof (intro exI conjI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
965 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
966 |
show "\<forall>i\<ge>n. eqlat (contract s p) ?c\<theta> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
967 |
proof (intro allI impI) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
968 |
fix i assume "n \<le> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
969 |
also have "i \<le> s i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
970 |
using increasing_inc by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
971 |
finally have "n \<le> s i" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
972 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
973 |
with fr have "is_fthread \<theta> p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
974 |
unfolding is_fthread_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
975 |
hence "has_fth p (s i) (s (Suc i)) (\<theta> (s i)) (\<theta> (s (Suc i)))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
976 |
unfolding has_fth_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
977 |
with less_imp_le[OF increasing_strict] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
978 |
have "eql (prod (p\<langle>s i,s (Suc i)\<rangle>)) (\<theta> (s i)) (\<theta> (s (Suc i)))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
979 |
by (simp add:Lemma7a) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
980 |
thus "eqlat (contract s p) ?c\<theta> i" unfolding contract_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
981 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
982 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
983 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
984 |
show "\<exists>\<^sub>\<infinity>i. descat (contract s p) ?c\<theta> i" |
27417 | 985 |
unfolding INFM_nat |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
986 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
987 |
fix i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
988 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
989 |
let ?K = "section_of s (max (s (Suc i)) n)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
990 |
from `\<exists>\<^sub>\<infinity>i. descat p \<theta> i` obtain j |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
991 |
where "s (Suc ?K) < j" "descat p \<theta> j" |
27417 | 992 |
unfolding INFM_nat by blast |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
993 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
994 |
let ?L = "section_of s j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
995 |
{ |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
996 |
fix x assume r: "x \<in> section s ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
997 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
998 |
have e1: "max (s (Suc i)) n < s (Suc ?K)" by (rule section_of2) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
999 |
note `s (Suc ?K) < j` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1000 |
also have "j < s (Suc ?L)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1001 |
by (rule section_of2) simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1002 |
finally have "Suc ?K \<le> ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1003 |
by (simp add:increasing_bij) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1004 |
with increasing_weak have "s (Suc ?K) \<le> s ?L" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1005 |
with e1 r have "max (s (Suc i)) n < x" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1006 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1007 |
hence "(s (Suc i)) < x" "n < x" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1008 |
} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1009 |
note range_est = this |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1010 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1011 |
have "is_desc_fthread \<theta> p (s ?L) (s (Suc ?L))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1012 |
unfolding is_desc_fthread_def is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1013 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1014 |
show "\<forall>m\<in>section s ?L. eqlat p \<theta> m" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1015 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1016 |
fix m assume "m\<in>section s ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1017 |
with range_est(2) have "n < m" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1018 |
with fr show "eqlat p \<theta> m" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1019 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1020 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1021 |
from in_section_of inc less_imp_le[OF `s (Suc ?K) < j`] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1022 |
have "j \<in> section s ?L" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1023 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1024 |
with `descat p \<theta> j` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1025 |
show "\<exists>m\<in>section s ?L. descat p \<theta> m" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1026 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1027 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1028 |
with less_imp_le[OF increasing_strict] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1029 |
have a: "descat (contract s p) ?c\<theta> ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1030 |
unfolding contract_def Lemma7b[symmetric] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1031 |
by (auto simp:Lemma7b[symmetric] has_desc_fth_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1032 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1033 |
have "i < ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1034 |
proof (rule classical) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1035 |
assume "\<not> i < ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1036 |
hence "s ?L < s (Suc i)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1037 |
by (simp add:increasing_bij) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1038 |
also have "\<dots> < s ?L" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1039 |
by (rule range_est(1)) (simp add:increasing_strict) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1040 |
finally show ?thesis . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1041 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1042 |
with a show "\<exists>l. i < l \<and> descat (contract s p) ?c\<theta> l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1043 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1044 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1045 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1046 |
with exI show "?B" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1047 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1048 |
assume "?B" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1049 |
then obtain \<theta> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1050 |
where dthread: "is_desc_thread \<theta> (contract s p)" .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1051 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1052 |
with dthreads_join inc |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1053 |
obtain \<theta>s where ths_spec: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1054 |
"desc (\<lambda>i. is_fthread (\<theta>s i) p (s i) (s (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1055 |
\<and> \<theta>s i (s i) = \<theta> i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1056 |
\<and> \<theta>s i (s (Suc i)) = \<theta> (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1057 |
(\<lambda>i. is_desc_fthread (\<theta>s i) p (s i) (s (Suc i)) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1058 |
\<and> \<theta>s i (s i) = \<theta> i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1059 |
\<and> \<theta>s i (s (Suc i)) = \<theta> (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1060 |
(is "desc ?alw ?inf") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1061 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1062 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1063 |
then obtain n where fr: "\<forall>i\<ge>n. ?alw i" by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1064 |
hence connected: "\<And>i. n < i \<Longrightarrow> \<theta>s i (s (Suc i)) = \<theta>s (Suc i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1065 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1066 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1067 |
let ?j\<theta> = "connect s \<theta>s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1068 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1069 |
from fr ths_spec have ths_spec2: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1070 |
"\<And>i. i > n \<Longrightarrow> is_fthread (\<theta>s i) p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1071 |
"\<exists>\<^sub>\<infinity>i. is_desc_fthread (\<theta>s i) p (s i) (s (Suc i))" |
27417 | 1072 |
by (auto intro:INFM_mono) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1073 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1074 |
have p1: "\<And>i. i > n \<Longrightarrow> is_fthread ?j\<theta> p (s i) (s (Suc i))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1075 |
by (rule connect_threads) (auto simp:connected ths_spec2) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1076 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1077 |
from ths_spec2(2) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1078 |
have "\<exists>\<^sub>\<infinity>i. n < i \<and> is_desc_fthread (\<theta>s i) p (s i) (s (Suc i))" |
28132 | 1079 |
unfolding INFM_drop_prefix . |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1080 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1081 |
hence p2: "\<exists>\<^sub>\<infinity>i. is_desc_fthread ?j\<theta> p (s i) (s (Suc i))" |
27417 | 1082 |
apply (rule INFM_mono) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1083 |
apply (rule connect_dthreads) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1084 |
by (auto simp:connected) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1085 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1086 |
with `increasing s` p1 |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1087 |
have "is_desc_thread ?j\<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1088 |
by (rule mk_inf_desc_thread) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1089 |
with exI show "?A" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1090 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1091 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1092 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1093 |
lemma repeated_edge: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1094 |
assumes "\<And>i. i > n \<Longrightarrow> dsc (snd (p i)) k k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1095 |
shows "is_desc_thread (\<lambda>i. k) p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1096 |
proof- |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1097 |
have th: "\<forall> m. \<exists>na>m. n < na" by arith |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1098 |
show ?thesis using prems |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1099 |
unfolding is_desc_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1100 |
apply (auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1101 |
apply (rule_tac x="Suc n" in exI, auto) |
27417 | 1102 |
apply (rule INFM_mono[where P="\<lambda>i. n < i"]) |
1103 |
apply (simp only:INFM_nat) |
|
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1104 |
by (auto simp add: th) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1105 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1106 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1107 |
lemma fin_from_inf: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1108 |
assumes "is_thread n \<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1109 |
assumes "n < i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1110 |
assumes "i < j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1111 |
shows "is_fthread \<theta> p i j" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1112 |
using prems |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1113 |
unfolding is_thread_def is_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1114 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1115 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1116 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1117 |
subsection {* Ramsey's Theorem *} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1118 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1119 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1120 |
"set2pair S = (THE (x,y). x < y \<and> S = {x,y})" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1121 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1122 |
lemma set2pair_conv: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1123 |
fixes x y :: nat |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1124 |
assumes "x < y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1125 |
shows "set2pair {x, y} = (x, y)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1126 |
unfolding set2pair_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1127 |
proof (rule the_equality, simp_all only:split_conv split_paired_all) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1128 |
from `x < y` show "x < y \<and> {x,y}={x,y}" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1129 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1130 |
fix a b |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1131 |
assume a: "a < b \<and> {x, y} = {a, b}" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1132 |
hence "{a, b} = {x, y}" by simp_all |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1133 |
hence "(a, b) = (x, y) \<or> (a, b) = (y, x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1134 |
by (cases "x = y") auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1135 |
thus "(a, b) = (x, y)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1136 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1137 |
assume "(a, b) = (y, x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1138 |
with a and `x < y` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1139 |
show ?thesis by auto (* contradiction *) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1140 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1141 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1142 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1143 |
definition |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1144 |
"set2list = inv set" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1145 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1146 |
lemma finite_set2list: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1147 |
assumes "finite S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1148 |
shows "set (set2list S) = S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1149 |
unfolding set2list_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1150 |
proof (rule f_inv_f) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1151 |
from `finite S` have "\<exists>l. set l = S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1152 |
by (rule finite_list) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1153 |
thus "S \<in> range set" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1154 |
unfolding image_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1155 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1156 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1157 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1158 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1159 |
corollary RamseyNatpairs: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1160 |
fixes S :: "'a set" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1161 |
and f :: "nat \<times> nat \<Rightarrow> 'a" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1162 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1163 |
assumes "finite S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1164 |
and inS: "\<And>x y. x < y \<Longrightarrow> f (x, y) \<in> S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1165 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1166 |
obtains T :: "nat set" and s :: "'a" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1167 |
where "infinite T" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1168 |
and "s \<in> S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1169 |
and "\<And>x y. \<lbrakk>x \<in> T; y \<in> T; x < y\<rbrakk> \<Longrightarrow> f (x, y) = s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1170 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1171 |
from `finite S` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1172 |
have "set (set2list S) = S" by (rule finite_set2list) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1173 |
then |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1174 |
obtain l where S: "S = set l" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1175 |
also from set_conv_nth have "\<dots> = {l ! i |i. i < length l}" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1176 |
finally have "S = {l ! i |i. i < length l}" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1177 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1178 |
let ?s = "length l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1179 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1180 |
from inS |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1181 |
have index_less: "\<And>x y. x \<noteq> y \<Longrightarrow> index_of l (f (set2pair {x, y})) < ?s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1182 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1183 |
fix x y :: nat |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1184 |
assume neq: "x \<noteq> y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1185 |
have "f (set2pair {x, y}) \<in> S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1186 |
proof (cases "x < y") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1187 |
case True hence "set2pair {x, y} = (x, y)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1188 |
by (rule set2pair_conv) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1189 |
with True inS |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1190 |
show ?thesis by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1191 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1192 |
case False |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1193 |
with neq have y_less: "y < x" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1194 |
have x:"{x,y} = {y,x}" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1195 |
with y_less have "set2pair {x, y} = (y, x)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1196 |
by (simp add:set2pair_conv) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1197 |
with y_less inS |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1198 |
show ?thesis by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1199 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1200 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1201 |
thus "index_of l (f (set2pair {x, y})) < length l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1202 |
by (simp add: S index_of_length) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1203 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1204 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1205 |
have "\<exists>Y. infinite Y \<and> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1206 |
(\<exists>t. t < ?s \<and> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1207 |
(\<forall>x\<in>Y. \<forall>y\<in>Y. x \<noteq> y \<longrightarrow> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1208 |
index_of l (f (set2pair {x, y})) = t))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1209 |
by (rule Ramsey2[of "UNIV::nat set", simplified]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1210 |
(auto simp:index_less) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1211 |
then obtain T i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1212 |
where inf: "infinite T" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1213 |
and i: "i < length l" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1214 |
and d: "\<And>x y. \<lbrakk>x \<in> T; y\<in>T; x \<noteq> y\<rbrakk> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1215 |
\<Longrightarrow> index_of l (f (set2pair {x, y})) = i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1216 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1217 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1218 |
have "l ! i \<in> S" unfolding S using i |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1219 |
by (rule nth_mem) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1220 |
moreover |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1221 |
have "\<And>x y. x \<in> T \<Longrightarrow> y\<in>T \<Longrightarrow> x < y |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1222 |
\<Longrightarrow> f (x, y) = l ! i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1223 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1224 |
fix x y assume "x \<in> T" "y \<in> T" "x < y" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1225 |
with d have |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1226 |
"index_of l (f (set2pair {x, y})) = i" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1227 |
with `x < y` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1228 |
have "i = index_of l (f (x, y))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1229 |
by (simp add:set2pair_conv) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1230 |
with `i < length l` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1231 |
show "f (x, y) = l ! i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1232 |
by (auto intro:index_of_member[symmetric] iff:index_of_length) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1233 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1234 |
moreover note inf |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1235 |
ultimately |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1236 |
show ?thesis using prems |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1237 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1238 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1239 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1240 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1241 |
subsection {* Main Result *} |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1242 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1243 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1244 |
theorem LJA_Theorem4: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1245 |
assumes "finite_acg A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1246 |
shows "SCT A \<longleftrightarrow> SCT' A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1247 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1248 |
assume "SCT A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1249 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1250 |
show "SCT' A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1251 |
proof (rule classical) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1252 |
assume "\<not> SCT' A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1253 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1254 |
then obtain n G |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1255 |
where in_closure: "(tcl A) \<turnstile> n \<leadsto>\<^bsup>G\<^esup> n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1256 |
and idemp: "G * G = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1257 |
and no_strict_arc: "\<forall>p. \<not>(G \<turnstile> p \<leadsto>\<^bsup>\<down>\<^esup> p)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1258 |
unfolding SCT'_def no_bad_graphs_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1259 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1260 |
from in_closure obtain k |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1261 |
where k_pow: "A ^ k \<turnstile> n \<leadsto>\<^bsup>G\<^esup> n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1262 |
and "0 < k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1263 |
unfolding in_tcl by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1264 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1265 |
from power_induces_path k_pow |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1266 |
obtain loop where loop_props: |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1267 |
"has_fpath A loop" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1268 |
"n = fst loop" "n = end_node loop" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1269 |
"G = prod loop" "k = length (snd loop)" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1270 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1271 |
with `0 < k` and path_loop_graph |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1272 |
have "has_ipath A (omega loop)" by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1273 |
with `SCT A` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1274 |
have thread: "\<exists>\<theta>. is_desc_thread \<theta> (omega loop)" by (auto simp:SCT_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1275 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1276 |
let ?s = "\<lambda>i. k * i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1277 |
let ?cp = "\<lambda>i::nat. (n, G)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1278 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1279 |
from loop_props have "fst loop = end_node loop" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1280 |
with `0 < k` `k = length (snd loop)` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1281 |
have "\<And>i. (omega loop)\<langle>?s i,?s (Suc i)\<rangle> = loop" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1282 |
by (rule sub_path_loop) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1283 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1284 |
with `n = fst loop` `G = prod loop` `k = length (snd loop)` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1285 |
have a: "contract ?s (omega loop) = ?cp" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1286 |
unfolding contract_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1287 |
by (simp add:path_loop_def split_def fst_p0) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1288 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1289 |
from `0 < k` have "increasing ?s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1290 |
by (auto simp:increasing_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1291 |
with thread have "\<exists>\<theta>. is_desc_thread \<theta> ?cp" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1292 |
unfolding a[symmetric] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1293 |
by (unfold contract_keeps_threads[symmetric]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1294 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1295 |
then obtain \<theta> where desc: "is_desc_thread \<theta> ?cp" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1296 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1297 |
then obtain n where thr: "is_thread n \<theta> ?cp" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1298 |
unfolding is_desc_thread_def is_thread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1299 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1300 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1301 |
have "finite (range \<theta>)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1302 |
proof (rule finite_range_ignore_prefix) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1303 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1304 |
from `finite_acg A` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1305 |
have "finite_acg (tcl A)" by (simp add:finite_tcl) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1306 |
with in_closure have "finite_graph G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1307 |
unfolding finite_acg_def all_finite_def by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1308 |
thus "finite (nodes G)" by (rule finite_nodes) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1309 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1310 |
from thread_image_nodes[OF thr] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1311 |
show "\<forall>i\<ge>n. \<theta> i \<in> nodes G" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1312 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1313 |
with finite_range |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1314 |
obtain p where inf_visit: "\<exists>\<^sub>\<infinity>i. \<theta> i = p" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1315 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1316 |
then obtain i where "n < i" "\<theta> i = p" |
27417 | 1317 |
by (auto simp:INFM_nat) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1318 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1319 |
from desc |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1320 |
have "\<exists>\<^sub>\<infinity>i. descat ?cp \<theta> i" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1321 |
unfolding is_desc_thread_def by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1322 |
then obtain j |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1323 |
where "i < j" and "descat ?cp \<theta> j" |
27417 | 1324 |
unfolding INFM_nat by auto |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1325 |
from inf_visit obtain k where "j < k" "\<theta> k = p" |
27417 | 1326 |
by (auto simp:INFM_nat) |
25314
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1327 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1328 |
from `i < j` `j < k` `n < i` thr |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1329 |
fin_from_inf[of n \<theta> ?cp] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1330 |
`descat ?cp \<theta> j` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1331 |
have "is_desc_fthread \<theta> ?cp i k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1332 |
unfolding is_desc_fthread_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1333 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1334 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1335 |
with `\<theta> k = p` `\<theta> i = p` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1336 |
have dfth: "has_desc_fth ?cp i k p p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1337 |
unfolding has_desc_fth_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1338 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1339 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1340 |
from `i < j` `j < k` have "i < k" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1341 |
hence "prod (?cp\<langle>i, k\<rangle>) = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1342 |
proof (induct i rule:strict_inc_induct) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1343 |
case base thus ?case by (simp add:sub_path_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1344 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1345 |
case (step i) thus ?case |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1346 |
by (simp add:sub_path_def upt_rec[of i k] idemp) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1347 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1348 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1349 |
with `i < j` `j < k` dfth Lemma7b[of i k ?cp p p] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1350 |
have "dsc G p p" by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1351 |
with no_strict_arc have False by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1352 |
thus ?thesis .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1353 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1354 |
next |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1355 |
assume "SCT' A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1356 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1357 |
show "SCT A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1358 |
proof (rule classical) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1359 |
assume "\<not> SCT A" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1360 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1361 |
with SCT_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1362 |
obtain p |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1363 |
where ipath: "has_ipath A p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1364 |
and no_desc_th: "\<not> (\<exists>\<theta>. is_desc_thread \<theta> p)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1365 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1366 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1367 |
from `finite_acg A` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1368 |
have "finite_acg (tcl A)" by (simp add: finite_tcl) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1369 |
hence "finite (dest_graph (tcl A))" (is "finite ?AG") |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1370 |
by (simp add: finite_acg_def finite_graph_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1371 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1372 |
from pdesc_acgplus[OF ipath] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1373 |
have a: "\<And>x y. x<y \<Longrightarrow> pdesc p\<langle>x,y\<rangle> \<in> dest_graph (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1374 |
unfolding has_edge_def . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1375 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1376 |
obtain S G |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1377 |
where "infinite S" "G \<in> dest_graph (tcl A)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1378 |
and all_G: "\<And>x y. \<lbrakk> x \<in> S; y \<in> S; x < y\<rbrakk> \<Longrightarrow> |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1379 |
pdesc (p\<langle>x,y\<rangle>) = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1380 |
apply (rule RamseyNatpairs[of ?AG "\<lambda>(x,y). pdesc p\<langle>x, y\<rangle>"]) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1381 |
apply (rule `finite ?AG`) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1382 |
by (simp only:split_conv, rule a, auto) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1383 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1384 |
obtain n H m where |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1385 |
G_struct: "G = (n, H, m)" by (cases G) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1386 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1387 |
let ?s = "enumerate S" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1388 |
let ?q = "contract ?s p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1389 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1390 |
note all_in_S[simp] = enumerate_in_set[OF `infinite S`] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1391 |
from `infinite S` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1392 |
have inc[simp]: "increasing ?s" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1393 |
unfolding increasing_def by (simp add:enumerate_mono) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1394 |
note increasing_bij[OF this, simp] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1395 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1396 |
from ipath_contract inc ipath |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1397 |
have "has_ipath (tcl A) ?q" . |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1398 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1399 |
from all_G G_struct |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1400 |
have all_H: "\<And>i. (snd (?q i)) = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1401 |
unfolding contract_def |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1402 |
by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1403 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1404 |
have loop: "(tcl A) \<turnstile> n \<leadsto>\<^bsup>H\<^esup> n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1405 |
and idemp: "H * H = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1406 |
proof - |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1407 |
let ?i = "?s 0" and ?j = "?s (Suc 0)" and ?k = "?s (Suc (Suc 0))" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1408 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1409 |
have "pdesc (p\<langle>?i,?j\<rangle>) = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1410 |
and "pdesc (p\<langle>?j,?k\<rangle>) = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1411 |
and "pdesc (p\<langle>?i,?k\<rangle>) = G" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1412 |
using all_G |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1413 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1414 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1415 |
with G_struct |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1416 |
have "m = end_node (p\<langle>?i,?j\<rangle>)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1417 |
"n = fst (p\<langle>?j,?k\<rangle>)" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1418 |
and Hs: "prod (p\<langle>?i,?j\<rangle>) = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1419 |
"prod (p\<langle>?j,?k\<rangle>) = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1420 |
"prod (p\<langle>?i,?k\<rangle>) = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1421 |
by auto |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1422 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1423 |
hence "m = n" by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1424 |
thus "tcl A \<turnstile> n \<leadsto>\<^bsup>H\<^esup> n" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1425 |
using G_struct `G \<in> dest_graph (tcl A)` |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1426 |
by (simp add:has_edge_def) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1427 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1428 |
from sub_path_prod[of ?i ?j ?k p] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1429 |
show "H * H = H" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1430 |
unfolding Hs by simp |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1431 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1432 |
moreover have "\<And>k. \<not>dsc H k k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1433 |
proof |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1434 |
fix k :: 'a assume "dsc H k k" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1435 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1436 |
with all_H repeated_edge |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1437 |
have "\<exists>\<theta>. is_desc_thread \<theta> ?q" by fast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1438 |
with inc have "\<exists>\<theta>. is_desc_thread \<theta> p" |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1439 |
by (subst contract_keeps_threads) |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1440 |
with no_desc_th |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1441 |
show False .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1442 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1443 |
ultimately |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1444 |
have False |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1445 |
using `SCT' A`[unfolded SCT'_def no_bad_graphs_def] |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1446 |
by blast |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1447 |
thus ?thesis .. |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1448 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1449 |
qed |
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1450 |
|
5eaf3e8b50a4
moved stuff about size change termination to its own session
krauss
parents:
diff
changeset
|
1451 |
end |