src/HOL/SizeChange/Implementation.thy
author nipkow
Fri, 13 Mar 2009 12:32:29 +0100
changeset 30502 b80d2621caee
parent 28562 4e74209f113e
child 31979 09f65e860bdb
permissions -rw-r--r--
hiding numeric coercions in LaTeX
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Library/SCT_Implementation.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 {* Implemtation of the SCT criterion *}
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 Implementation
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     9
imports Correctness
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
fun edges_match :: "('n \<times> 'e \<times> 'n) \<times> ('n \<times> 'e \<times> 'n) \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    13
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    14
  "edges_match ((n, e, m), (n',e',m')) = (m = n')"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
fun connect_edges :: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    17
  "('n \<times> ('e::times) \<times> 'n) \<times> ('n \<times> 'e \<times> 'n)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    18
  \<Rightarrow> ('n \<times> 'e \<times> 'n)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    19
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
  "connect_edges ((n,e,m), (n', e', m')) = (n, e * e', m')"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    21
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    22
lemma grcomp_code [code]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
  "grcomp (Graph G) (Graph H) = Graph (connect_edges ` { x \<in> G\<times>H. edges_match x })"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    24
  by (rule graph_ext) (auto simp:graph_mult_def has_edge_def image_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    25
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    26
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    27
lemma mk_tcl_finite_terminates:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
  fixes A :: "'a acg"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
  assumes fA: "finite_acg A" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
  shows "mk_tcl_dom (A, A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
  from fA have fin_tcl: "finite_acg (tcl A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
    by (simp add:finite_tcl)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    35
  hence "finite (dest_graph (tcl A))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    36
    unfolding finite_acg_def finite_graph_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
  let ?count = "\<lambda>G. card (dest_graph G)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
  let ?N = "?count (tcl A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
  let ?m = "\<lambda>X. ?N - (?count X)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
  let ?P = "\<lambda>X. mk_tcl_dom (A, X)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    43
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    44
  {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
    fix X
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
    assume "X \<le> tcl A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    47
    then
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
    have "mk_tcl_dom (A, X)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
    proof (induct X rule:measure_induct_rule[of ?m])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
      case (less X)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    51
      show ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    52
      proof (cases "X * A \<le> X")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
        case True 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    54
        with mk_tcl.domintros show ?thesis by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    55
      next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    56
        case False
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    57
        then have l: "X < X + X * A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    58
          unfolding graph_less_def graph_leq_def graph_plus_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
          by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    61
        from `X \<le> tcl A` 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    62
        have "X * A \<le> tcl A * A" by (simp add:mult_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    63
        also have "\<dots> \<le> A + tcl A * A" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    64
        also have "\<dots> = tcl A" by (simp add:tcl_unfold_right[symmetric])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    65
        finally have "X * A \<le> tcl A" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    66
        with `X \<le> tcl A` 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    67
        have "X + X * A \<le> tcl A + tcl A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    68
          by (rule add_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    69
        hence less_tcl: "X + X * A \<le> tcl A" by simp 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    70
        hence "X < tcl A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    71
          using l `X \<le> tcl A` by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    72
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    73
        from less_tcl fin_tcl
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
        have "finite_acg (X + X * A)" by (rule finite_acg_subset)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
        hence "finite (dest_graph (X + X * A))" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
          unfolding finite_acg_def finite_graph_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    77
        
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
        hence X: "?count X < ?count (X + X * A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
          using l[simplified graph_less_def graph_leq_def]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    80
          by (rule psubset_card_mono)
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
        have "?count X < ?N" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    83
          apply (rule psubset_card_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    84
          by fact (rule `X < tcl A`[simplified graph_less_def])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    85
        
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    86
        with X have "?m (X + X * A) < ?m X" by arith
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    87
        
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    88
        from  less.hyps this less_tcl
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    89
        have "mk_tcl_dom (A, X + X * A)" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    90
        with mk_tcl.domintros show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    91
      qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    92
    qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    93
  }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    94
  from this less_tcl show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    95
qed
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
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    98
lemma mk_tcl_finite_tcl:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    99
  fixes A :: "'a acg"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   100
  assumes fA: "finite_acg A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   101
  shows "mk_tcl A A = tcl A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   102
  using mk_tcl_finite_terminates[OF fA]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   103
  by (simp only: tcl_def mk_tcl_correctness star_commute)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   104
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   105
definition test_SCT :: "nat acg \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   106
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   107
  "test_SCT \<A> = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   108
  (let \<T> = mk_tcl \<A> \<A>
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   109
    in (\<forall>(n,G,m)\<in>dest_graph \<T>. 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   110
          n \<noteq> m \<or> G * G \<noteq> G \<or> 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   111
         (\<exists>(p::nat,e,q)\<in>dest_graph G. p = q \<and> e = LESS)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   112
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   113
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   114
lemma SCT'_exec:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   115
  assumes fin: "finite_acg A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   116
  shows "SCT' A = test_SCT A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   117
  using mk_tcl_finite_tcl[OF fin]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   118
  unfolding test_SCT_def Let_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   119
  unfolding SCT'_def no_bad_graphs_def has_edge_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   120
  by force
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   121
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   122
code_modulename SML
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   123
  Implementation Graphs
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   124
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27436
diff changeset
   125
lemma [code]:
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   126
  "(G\<Colon>('a\<Colon>eq, 'b\<Colon>eq) graph) \<le> H \<longleftrightarrow> dest_graph G \<subseteq> dest_graph H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   127
  "(G\<Colon>('a\<Colon>eq, 'b\<Colon>eq) graph) < H \<longleftrightarrow> dest_graph G \<subset> dest_graph H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   128
  unfolding graph_leq_def graph_less_def by rule+
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   129
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27436
diff changeset
   130
lemma [code]:
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   131
  "(G\<Colon>('a\<Colon>eq, 'b\<Colon>eq) graph) + H = Graph (dest_graph G \<union> dest_graph H)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   132
  unfolding graph_plus_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   133
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27436
diff changeset
   134
lemma [code]:
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   135
  "(G\<Colon>('a\<Colon>eq, 'b\<Colon>{eq, times}) graph) * H = grcomp G H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   136
  unfolding graph_mult_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   137
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   138
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   139
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   140
lemma SCT'_empty: "SCT' (Graph {})"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   141
  unfolding SCT'_def no_bad_graphs_def graph_zero_def[symmetric]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   142
  tcl_zero
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   143
  by (simp add:in_grzero)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   144
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   145
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   146
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   147
subsection {* Witness checking *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   148
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   149
definition test_SCT_witness :: "nat acg \<Rightarrow> nat acg \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   150
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   151
  "test_SCT_witness A T = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   152
  (A \<le> T \<and> A * T \<le> T \<and>
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   153
       (\<forall>(n,G,m)\<in>dest_graph T. 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   154
          n \<noteq> m \<or> G * G \<noteq> G \<or> 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   155
         (\<exists>(p::nat,e,q)\<in>dest_graph G. p = q \<and> e = LESS)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   156
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   157
lemma no_bad_graphs_ucl:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   158
  assumes "A \<le> B"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   159
  assumes "no_bad_graphs B"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   160
  shows "no_bad_graphs A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   161
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   162
  unfolding no_bad_graphs_def has_edge_def graph_leq_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   163
  by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   164
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   165
lemma SCT'_witness:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   166
  assumes a: "test_SCT_witness A T"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   167
  shows "SCT' A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   168
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   169
  from a have "A \<le> T" "A * T \<le> T" by (auto simp:test_SCT_witness_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   170
  hence "A + A * T \<le> T" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   171
    by (subst add_idem[of T, symmetric], rule add_mono)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   172
  with star3' have "tcl A \<le> T" unfolding tcl_def .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   173
  moreover
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   174
  from a have "no_bad_graphs T"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   175
    unfolding no_bad_graphs_def test_SCT_witness_def has_edge_def
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
  ultimately
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   178
  show ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   179
    unfolding SCT'_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   180
    by (rule no_bad_graphs_ucl)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   181
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   182
27436
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 26822
diff changeset
   183
(* ML {* @{code test_SCT} *} *)
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   184
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   185
end