src/HOL/SizeChange/Graphs.thy
author ballarin
Fri, 19 Dec 2008 16:39:23 +0100
changeset 29249 4dc278c8dc59
parent 28685 275122631271
child 30282 a16af775e08d
permissions -rw-r--r--
All logics ported to new locales.
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/Graphs.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 {* General Graphs as Sets *}
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 Graphs
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     9
imports Main Misc_Tools Kleene_Algebras
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 {* Basic types, Size Change Graphs *}
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
datatype ('a, 'b) graph = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
  Graph "('a \<times> 'b \<times> 'a) set"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    17
primrec dest_graph :: "('a, 'b) graph \<Rightarrow> ('a \<times> 'b \<times> 'a) set"
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    18
  where "dest_graph (Graph G) = G"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    19
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
lemma graph_dest_graph[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    21
  "Graph (dest_graph G) = G"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    22
  by (cases G) simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    24
lemma split_graph_all:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    25
  "(\<And>gr. PROP P gr) \<equiv> (\<And>set. PROP P (Graph set))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    26
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    27
  fix set
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
  assume "\<And>gr. PROP P gr"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
  then show "PROP P (Graph set)" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
  fix gr
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
  assume "\<And>set. PROP P (Graph set)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
  then have "PROP P (Graph (dest_graph gr))" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
  then show "PROP P gr" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    35
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    36
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
definition 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
  has_edge :: "('n,'e) graph \<Rightarrow> 'n \<Rightarrow> 'e \<Rightarrow> 'n \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
("_ \<turnstile> _ \<leadsto>\<^bsup>_\<^esup> _")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
  "has_edge G n e n' = ((n, e, n') \<in> dest_graph G)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
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
subsection {* Graph composition *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
fun grcomp :: "('n, 'e::times) graph \<Rightarrow> ('n, 'e) graph  \<Rightarrow> ('n, 'e) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    47
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
  "grcomp (Graph G) (Graph H) = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
  Graph {(p,b,q) | p b q. 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
  (\<exists>k e e'. (p,e,k)\<in>G \<and> (k,e',q)\<in>H \<and> b = e * e')}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    51
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    52
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
declare grcomp.simps[code del]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    54
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    55
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    56
lemma graph_ext:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    57
  assumes "\<And>n e n'. has_edge G n e n' = has_edge H n e n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    58
  shows "G = H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
  by (cases G, cases H) (auto simp:split_paired_all has_edge_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    61
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    62
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    63
instantiation graph :: (type, type) comm_monoid_add
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    64
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    65
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    66
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    67
  graph_zero_def: "0 = Graph {}" 
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    68
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    69
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
    70
  graph_plus_def [code del]: "G + H = Graph (dest_graph G \<union> dest_graph H)"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    71
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    72
instance proof
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    73
  fix x y z :: "('a,'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
  show "x + y + z = x + (y + z)" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
   and "x + y = y + x" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
   and "0 + x = x"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    77
  unfolding graph_plus_def graph_zero_def by auto
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    80
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    81
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    82
instantiation graph :: (type, type) "{distrib_lattice, complete_lattice}"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    83
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    84
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    85
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
    86
  graph_leq_def [code del]: "G \<le> H \<longleftrightarrow> dest_graph G \<subseteq> dest_graph H"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    87
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    88
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
    89
  graph_less_def [code del]: "G < H \<longleftrightarrow> dest_graph G \<subset> dest_graph H"
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    90
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    91
definition
28685
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    92
  [code del]: "bot = Graph {}"
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    93
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    94
definition
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    95
  [code del]: "top = Graph UNIV"
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    96
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
    97
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
    98
  [code del]: "inf G H = Graph (dest_graph G \<inter> dest_graph H)"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    99
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   100
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
   101
  [code del]: "sup (G \<Colon> ('a, 'b) graph) H = G + H"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   102
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   103
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
   104
  Inf_graph_def [code del]: "Inf = (\<lambda>Gs. Graph (\<Inter>(dest_graph ` Gs)))"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   105
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   106
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
   107
  Sup_graph_def [code del]: "Sup = (\<lambda>Gs. Graph (\<Union>(dest_graph ` Gs)))"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   108
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   109
instance proof
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   110
  fix x y z :: "('a,'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   111
  fix A :: "('a, 'b) graph set"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   112
27682
25aceefd4786 added class preorder
haftmann
parents: 25764
diff changeset
   113
  show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   114
    unfolding graph_leq_def graph_less_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   115
    by (cases x, cases y) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   116
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   117
  show "x \<le> x" unfolding graph_leq_def ..
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   118
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   119
  { assume "x \<le> y" "y \<le> z" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   120
    with order_trans show "x \<le> z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   121
      unfolding graph_leq_def . }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   122
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   123
  { assume "x \<le> y" "y \<le> x" thus "x = y" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   124
      unfolding graph_leq_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   125
      by (cases x, cases y) simp }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   126
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   127
  show "inf x y \<le> x" "inf x y \<le> y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   128
    unfolding inf_graph_def graph_leq_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   129
    by auto    
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   130
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   131
  { assume "x \<le> y" "x \<le> z" thus "x \<le> inf y z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   132
      unfolding inf_graph_def graph_leq_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   133
      by auto }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   134
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   135
  show "x \<le> sup x y" "y \<le> sup x y"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   136
    unfolding sup_graph_def graph_leq_def graph_plus_def by auto
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
  { assume "y \<le> x" "z \<le> x" thus "sup y z \<le> x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   139
      unfolding sup_graph_def graph_leq_def graph_plus_def by auto }
28685
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
   140
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
   141
  show "bot \<le> x" unfolding graph_leq_def bot_graph_def by simp
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
   142
275122631271 new classes "top" and "bot"
haftmann
parents: 28562
diff changeset
   143
  show "x \<le> top" unfolding graph_leq_def top_graph_def by simp
25314
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
  show "sup x (inf y z) = inf (sup x y) (sup x z)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   146
    unfolding inf_graph_def sup_graph_def graph_leq_def graph_plus_def 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
  { assume "x \<in> A" thus "Inf A \<le> x" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   149
      unfolding Inf_graph_def graph_leq_def by auto }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   150
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   151
  { assume "\<And>x. x \<in> A \<Longrightarrow> z \<le> x" thus "z \<le> Inf A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   152
    unfolding Inf_graph_def graph_leq_def by auto }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   153
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   154
  { assume "x \<in> A" thus "x \<le> Sup A" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   155
      unfolding Sup_graph_def graph_leq_def by auto }
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
  { assume "\<And>x. x \<in> A \<Longrightarrow> x \<le> z" thus "Sup A \<le> z"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   158
    unfolding Sup_graph_def graph_leq_def by auto }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   159
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   160
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   161
end
25314
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 in_grplus:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   164
  "has_edge (G + H) p b q = (has_edge G p b q \<or> has_edge H p b q)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   165
  by (cases G, cases H, auto simp:has_edge_def graph_plus_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   166
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   167
lemma in_grzero:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   168
  "has_edge 0 p b q = False"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   169
  by (simp add:graph_zero_def has_edge_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   170
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   171
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   172
subsubsection {* Multiplicative Structure *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   173
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   174
instantiation graph :: (type, times) mult_zero
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   175
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   176
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   177
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27682
diff changeset
   178
  graph_mult_def [code del]: "G * H = grcomp G H" 
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   179
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   180
instance proof
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   181
  fix a :: "('a, 'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   182
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   183
  show "0 * a = 0" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   184
    unfolding graph_mult_def graph_zero_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   185
    by (cases a) (simp add:grcomp.simps)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   186
  show "a * 0 = 0" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   187
    unfolding graph_mult_def graph_zero_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   188
    by (cases a) (simp add:grcomp.simps)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   189
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   190
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   191
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   192
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   193
instantiation graph :: (type, one) one 
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   194
begin
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   195
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   196
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   197
  graph_one_def: "1 = Graph { (x, 1, x) |x. True}"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   198
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   199
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   200
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   201
end
25314
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
lemma in_grcomp:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   204
  "has_edge (G * H) p b q
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   205
  = (\<exists>k e e'. has_edge G p e k \<and> has_edge H k e' q \<and> b = e * e')"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   206
  by (cases G, cases H) (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
   207
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   208
lemma in_grunit:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   209
  "has_edge 1 p b q = (p = q \<and> b = 1)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   210
  by (auto simp:graph_one_def has_edge_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   211
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   212
instance graph :: (type, semigroup_mult) semigroup_mult
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   213
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   214
  fix G1 G2 G3 :: "('a,'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   215
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   216
  show "G1 * G2 * G3 = G1 * (G2 * G3)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   217
  proof (rule graph_ext, rule trans)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   218
    fix p J q
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   219
    show "has_edge ((G1 * G2) * G3) p J q =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   220
      (\<exists>G i H j I.
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   221
      has_edge G1 p G i
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   222
      \<and> has_edge G2 i H j
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   223
      \<and> has_edge G3 j I q
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   224
      \<and> J = (G * H) * I)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   225
      by (simp only:in_grcomp) blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   226
    show "\<dots> = has_edge (G1 * (G2 * G3)) p J q"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   227
      by (simp only:in_grcomp mult_assoc) blast
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
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   231
instantiation graph :: (type, monoid_mult) "{semiring_1, idem_add, recpower, star}"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   232
begin
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   233
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   234
primrec power_graph :: "('a\<Colon>type, 'b\<Colon>monoid_mult) graph \<Rightarrow> nat => ('a, 'b) graph"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   235
where
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   236
  "(A \<Colon> ('a, 'b) graph) ^ 0 = 1"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   237
| "(A \<Colon> ('a, 'b) graph) ^ Suc n = A * (A ^ n)"
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   238
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   239
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   240
  graph_star_def: "star (G \<Colon> ('a, 'b) graph) = (SUP n. G ^ n)" 
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   241
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   242
instance proof
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   243
  fix a b c :: "('a, 'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   244
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   245
  show "1 * a = a" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   246
    by (rule graph_ext) (auto simp:in_grcomp in_grunit)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   247
  show "a * 1 = a"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   248
    by (rule graph_ext) (auto simp:in_grcomp in_grunit)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   249
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   250
  show "(a + b) * c = a * c + b * c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   251
    by (rule graph_ext, simp add:in_grcomp in_grplus) blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   252
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   253
  show "a * (b + c) = a * b + a * c"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   254
    by (rule graph_ext, simp add:in_grcomp in_grplus) blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   255
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   256
  show "(0::('a,'b) graph) \<noteq> 1" unfolding graph_zero_def graph_one_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   257
    by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   258
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   259
  show "a + a = a" unfolding graph_plus_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   260
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   261
  show "a ^ 0 = 1" "\<And>n. a ^ (Suc n) = a * a ^ n"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   262
    by simp_all
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   263
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   264
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   265
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   266
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   267
lemma graph_leqI:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   268
  assumes "\<And>n e n'. has_edge G n e n' \<Longrightarrow> has_edge H n e n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   269
  shows "G \<le> H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   270
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   271
  unfolding graph_leq_def has_edge_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   272
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   273
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   274
lemma in_graph_plusE:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   275
  assumes "has_edge (G + H) n e n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   276
  assumes "has_edge G n e n' \<Longrightarrow> P"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   277
  assumes "has_edge H n e n' \<Longrightarrow> P"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   278
  shows P
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   279
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   280
  by (auto simp: in_grplus)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   281
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   282
lemma in_graph_compE:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   283
  assumes GH: "has_edge (G * H) n e n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   284
  obtains e1 k e2 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   285
  where "has_edge G n e1 k" "has_edge H k e2 n'" "e = e1 * e2"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   286
  using GH
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   287
  by (auto simp: in_grcomp)
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 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   290
  assumes "x \<in> S k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   291
  shows "x \<in> (\<Union>k. S k)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   292
  using assms by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   293
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   294
lemma graph_union_least:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   295
  assumes "\<And>n. Graph (G n) \<le> C"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   296
  shows "Graph (\<Union>n. G n) \<le> C"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   297
  using assms unfolding graph_leq_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   298
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   299
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   300
lemma Sup_graph_eq:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   301
  "(SUP n. Graph (G n)) = Graph (\<Union>n. G n)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   302
proof (rule order_antisym)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   303
  show "(SUP n. Graph (G n)) \<le> Graph (\<Union>n. G n)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   304
    by  (rule SUP_leI) (auto simp add: graph_leq_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   305
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   306
  show "Graph (\<Union>n. G n) \<le> (SUP n. Graph (G n))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   307
  by (rule graph_union_least, rule le_SUPI', rule) 
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 has_edge_leq: "has_edge G p b q = (Graph {(p,b,q)} \<le> G)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   311
  unfolding has_edge_def graph_leq_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   312
  by (cases G) simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   313
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   315
lemma Sup_graph_eq2:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   316
  "(SUP n. G n) = Graph (\<Union>n. dest_graph (G n))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   317
  using Sup_graph_eq[of "\<lambda>n. dest_graph (G n)", simplified]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   318
  by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   319
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   320
lemma in_SUP:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   321
  "has_edge (SUP x. Gs x) p b q = (\<exists>x. has_edge (Gs x) p b q)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   322
  unfolding Sup_graph_eq2 has_edge_leq graph_leq_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   323
  by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   324
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   325
instance graph :: (type, monoid_mult) kleene_by_complete_lattice
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   326
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   327
  fix a b c :: "('a, 'b) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   328
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   329
  show "a \<le> b \<longleftrightarrow> a + b = b" unfolding graph_leq_def graph_plus_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   330
    by (cases a, cases b) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   331
27682
25aceefd4786 added class preorder
haftmann
parents: 25764
diff changeset
   332
  from less_le_not_le show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a" .
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   333
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   334
  show "a * star b * c = (SUP n. a * b ^ n * c)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   335
    unfolding graph_star_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   336
    by (rule graph_ext) (force simp:in_SUP in_grcomp)
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
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   340
lemma in_star: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   341
  "has_edge (star G) a x b = (\<exists>n. has_edge (G ^ n) a x b)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   342
  by (auto simp:graph_star_def in_SUP)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   343
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   344
lemma tcl_is_SUP:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   345
  "tcl (G::('a::type, 'b::monoid_mult) graph) =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   346
  (SUP n. G ^ (Suc n))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   347
  unfolding tcl_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   348
  using star_cont[of 1 G G]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   349
  by (simp add:power_Suc power_commutes)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   350
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   351
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   352
lemma in_tcl: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   353
  "has_edge (tcl G) a x b = (\<exists>n>0. has_edge (G ^ n) a x b)"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
   354
  apply (auto simp: tcl_is_SUP in_SUP simp del: power_graph.simps)
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   355
  apply (rule_tac x = "n - 1" in exI, auto)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   356
  done
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   357
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   358
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   359
subsection {* Infinite Paths *}
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
types ('n, 'e) ipath = "('n \<times> 'e) sequence"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   362
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   363
definition has_ipath :: "('n, 'e) graph \<Rightarrow> ('n, 'e) ipath \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   364
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   365
  "has_ipath G p = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   366
  (\<forall>i. has_edge G (fst (p i)) (snd (p i)) (fst (p (Suc i))))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   367
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
subsection {* Finite Paths *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   370
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   371
types ('n, 'e) fpath = "('n \<times> ('e \<times> 'n) list)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   372
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   373
inductive  has_fpath :: "('n, 'e) graph \<Rightarrow> ('n, 'e) fpath \<Rightarrow> bool" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   374
  for G :: "('n, 'e) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   375
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   376
  has_fpath_empty: "has_fpath G (n, [])"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   377
| has_fpath_join: "\<lbrakk>G \<turnstile> n \<leadsto>\<^bsup>e\<^esup> n'; has_fpath G (n', es)\<rbrakk> \<Longrightarrow> has_fpath G (n, (e, n')#es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   378
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   379
definition 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   380
  "end_node p = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   381
  (if snd p = [] then fst p else snd (snd p ! (length (snd p) - 1)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   382
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   383
definition path_nth :: "('n, 'e) fpath \<Rightarrow> nat \<Rightarrow> ('n \<times> 'e \<times> 'n)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   384
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   385
  "path_nth p k = (if k = 0 then fst p else snd (snd p ! (k - 1)), snd p ! k)"
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 endnode_nth:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   388
  assumes "length (snd p) = Suc k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   389
  shows "end_node p = snd (snd (path_nth p k))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   390
  using assms unfolding end_node_def path_nth_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   391
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   392
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   393
lemma path_nth_graph:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   394
  assumes "k < length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   395
  assumes "has_fpath G p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   396
  shows "(\<lambda>(n,e,n'). has_edge G n e n') (path_nth p k)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   397
using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   398
proof (induct k arbitrary: p)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   399
  case 0 thus ?case 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   400
    unfolding path_nth_def by (auto elim:has_fpath.cases)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   401
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   402
  case (Suc k p)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   403
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   404
  from `has_fpath G p` show ?case 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   405
  proof (rule has_fpath.cases)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   406
    case goal1 with Suc show ?case by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   407
  next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   408
    fix n e n' es
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   409
    assume st: "p = (n, (e, n') # es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   410
       "G \<turnstile> n \<leadsto>\<^bsup>e\<^esup> n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   411
       "has_fpath G (n', es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   412
    with Suc
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   413
    have "(\<lambda>(n, b, a). G \<turnstile> n \<leadsto>\<^bsup>b\<^esup> a) (path_nth (n', es) k)" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   414
    with st show ?thesis by (cases k, auto simp:path_nth_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   415
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   416
qed
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 path_nth_connected:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   419
  assumes "Suc k < length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   420
  shows "fst (path_nth p (Suc k)) = snd (snd (path_nth p k))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   421
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   422
  unfolding path_nth_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   423
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   424
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   425
definition path_loop :: "('n, 'e) fpath \<Rightarrow> ('n, 'e) ipath" ("omega")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   426
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   427
  "omega p \<equiv> (\<lambda>i. (\<lambda>(n,e,n'). (n,e)) (path_nth p (i mod (length (snd p)))))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   428
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   429
lemma fst_p0: "fst (path_nth p 0) = fst p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   430
  unfolding path_nth_def by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   431
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   432
lemma path_loop_connect:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   433
  assumes "fst p = end_node p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   434
  and "0 < length (snd p)" (is "0 < ?l")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   435
  shows "fst (path_nth p (Suc i mod (length (snd p))))
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   436
  = snd (snd (path_nth p (i mod length (snd p))))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   437
  (is "\<dots> = snd (snd (path_nth p ?k))")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   438
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   439
  from `0 < ?l` have "i mod ?l < ?l" (is "?k < ?l")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   440
    by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   441
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   442
  show ?thesis 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   443
  proof (cases "Suc ?k < ?l")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   444
    case True
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   445
    hence "Suc ?k \<noteq> ?l" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   446
    with path_nth_connected[OF True]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   447
    show ?thesis
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   448
      by (simp add:mod_Suc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   449
  next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   450
    case False 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   451
    with `?k < ?l` have wrap: "Suc ?k = ?l" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   452
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   453
    hence "fst (path_nth p (Suc i mod ?l)) = fst (path_nth p 0)" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   454
      by (simp add: mod_Suc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   455
    also from fst_p0 have "\<dots> = fst p" .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   456
    also have "\<dots> = end_node p" by fact
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   457
    also have "\<dots> = snd (snd (path_nth p ?k))" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   458
      by (auto simp: endnode_nth wrap)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   459
    finally show ?thesis .
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   460
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   461
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   462
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   463
lemma path_loop_graph:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   464
  assumes "has_fpath G p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   465
  and loop: "fst p = end_node p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   466
  and nonempty: "0 < length (snd p)" (is "0 < ?l")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   467
  shows "has_ipath G (omega p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   468
proof -
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   469
  {
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   470
    fix i 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   471
    from `0 < ?l` have "i mod ?l < ?l" (is "?k < ?l")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   472
      by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   473
    from this and `has_fpath G p`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   474
    have pk_G: "(\<lambda>(n,e,n'). has_edge G n e n') (path_nth p ?k)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   475
      by (rule path_nth_graph)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   476
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   477
    from path_loop_connect[OF loop nonempty] pk_G
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   478
    have "has_edge G (fst (omega p i)) (snd (omega p i)) (fst (omega p (Suc i)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   479
      unfolding path_loop_def has_edge_def split_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   480
      by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   481
  }
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   482
  then show ?thesis by (auto simp:has_ipath_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   483
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   484
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   485
definition prod :: "('n, 'e::monoid_mult) fpath \<Rightarrow> 'e"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   486
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   487
  "prod p = foldr (op *) (map fst (snd p)) 1"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   488
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   489
lemma prod_simps[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   490
  "prod (n, []) = 1"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   491
  "prod (n, (e,n')#es) = e * (prod (n',es))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   492
unfolding prod_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   493
by simp_all
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   494
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   495
lemma power_induces_path:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   496
  assumes a: "has_edge (A ^ k) n G m"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   497
  obtains p 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   498
    where "has_fpath A p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   499
      and "n = fst p" "m = end_node p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   500
      and "G = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   501
      and "k = length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   502
  using a
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   503
proof (induct k arbitrary:m n G thesis)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   504
  case (0 m n G)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   505
  let ?p = "(n, [])"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   506
  from 0 have "has_fpath A ?p" "m = end_node ?p" "G = prod ?p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   507
    by (auto simp:in_grunit end_node_def intro:has_fpath.intros)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   508
  thus ?case using 0 by (auto simp:end_node_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   509
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   510
  case (Suc k m n G)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   511
  hence "has_edge (A * A ^ k) n G m" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   512
    by (simp add:power_Suc power_commutes)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   513
  then obtain G' H j where 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   514
    a_A: "has_edge A n G' j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   515
    and H_pow: "has_edge (A ^ k) j H m"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   516
    and [simp]: "G = G' * H"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   517
    by (auto simp:in_grcomp) 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   518
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   519
  from H_pow and Suc
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   520
  obtain p
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   521
    where p_path: "has_fpath A p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   522
    and [simp]: "j = fst p" "m = end_node p" "H = prod p" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   523
    "k = length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   524
    by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   525
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   526
  let ?p' = "(n, (G', j)#snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   527
  from a_A and p_path
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   528
  have "has_fpath A ?p'" "m = end_node ?p'" "G = prod ?p'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   529
    by (auto simp:end_node_def nth.simps intro:has_fpath.intros split:nat.split)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   530
  thus ?case using Suc by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   531
qed
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
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   534
subsection {* Sub-Paths *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   535
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   536
definition sub_path :: "('n, 'e) ipath \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ('n, 'e) fpath"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   537
("(_\<langle>_,_\<rangle>)")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   538
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   539
  "p\<langle>i,j\<rangle> =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   540
  (fst (p i), map (\<lambda>k. (snd (p k), fst (p (Suc k)))) [i ..< j])"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   541
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   542
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   543
lemma sub_path_is_path: 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   544
  assumes ipath: "has_ipath G p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   545
  assumes l: "i \<le> j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   546
  shows "has_fpath G (p\<langle>i,j\<rangle>)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   547
  using l
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   548
proof (induct i rule:inc_induct)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   549
  case base show ?case by (auto simp:sub_path_def intro:has_fpath.intros)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   550
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   551
  case (step i)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   552
  with ipath upt_rec[of i j]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   553
  show ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   554
    by (auto simp:sub_path_def has_ipath_def intro:has_fpath.intros)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   555
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   556
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
lemma sub_path_start[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   559
  "fst (p\<langle>i,j\<rangle>) = fst (p i)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   560
  by (simp add:sub_path_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   561
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   562
lemma nth_upto[simp]: "k < j - i \<Longrightarrow> [i ..< j] ! k = i + k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   563
  by (induct k) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   564
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   565
lemma sub_path_end[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   566
  "i < j \<Longrightarrow> end_node (p\<langle>i,j\<rangle>) = fst (p j)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   567
  by (auto simp:sub_path_def end_node_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   568
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   569
lemma foldr_map: "foldr f (map g xs) = foldr (f o g) xs"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   570
  by (induct xs) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   571
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   572
lemma upto_append[simp]:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   573
  assumes "i \<le> j" "j \<le> k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   574
  shows "[ i ..< j ] @ [j ..< k] = [i ..< k]"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   575
  using assms and upt_add_eq_append[of i j "k - j"]
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   576
  by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   577
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   578
lemma foldr_monoid: "foldr (op *) xs 1 * foldr (op *) ys 1
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   579
  = foldr (op *) (xs @ ys) (1::'a::monoid_mult)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   580
  by (induct xs) (auto simp:mult_assoc)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   581
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   582
lemma sub_path_prod:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   583
  assumes "i < j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   584
  assumes "j < k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   585
  shows "prod (p\<langle>i,k\<rangle>) = prod (p\<langle>i,j\<rangle>) * prod (p\<langle>j,k\<rangle>)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   586
  using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   587
  unfolding prod_def sub_path_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   588
  by (simp add:map_compose[symmetric] comp_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   589
   (simp only:foldr_monoid map_append[symmetric] upto_append)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   590
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   591
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   592
lemma path_acgpow_aux:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   593
  assumes "length es = l"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   594
  assumes "has_fpath G (n,es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   595
  shows "has_edge (G ^ l) n (prod (n,es)) (end_node (n,es))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   596
using assms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   597
proof (induct l arbitrary:n es)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   598
  case 0 thus ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   599
    by (simp add:in_grunit end_node_def) 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   600
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   601
  case (Suc l n es)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   602
  hence "es \<noteq> []" by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   603
  let ?n' = "snd (hd es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   604
  let ?es' = "tl es"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   605
  let ?e = "fst (hd es)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   606
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   607
  from Suc have len: "length ?es' = l" by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   608
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   609
  from Suc
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   610
  have [simp]: "end_node (n, es) = end_node (?n', ?es')"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   611
    by (cases es) (auto simp:end_node_def nth.simps split:nat.split)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   612
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   613
  from `has_fpath G (n,es)`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   614
  have "has_fpath G (?n', ?es')"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   615
    by (rule has_fpath.cases) (auto intro:has_fpath.intros)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   616
  with Suc len
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   617
  have "has_edge (G ^ l) ?n' (prod (?n', ?es')) (end_node (?n', ?es'))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   618
    by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   619
  moreover
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   620
  from `es \<noteq> []`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   621
  have "prod (n, es) = ?e * (prod (?n', ?es'))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   622
    by (cases es) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   623
  moreover
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   624
  from `has_fpath G (n,es)` have c:"has_edge G n ?e ?n'"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   625
    by (rule has_fpath.cases) (insert `es \<noteq> []`, auto)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   626
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   627
  ultimately
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   628
  show ?case
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   629
     unfolding power_Suc 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   630
     by (auto simp:in_grcomp)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   631
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   632
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   633
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   634
lemma path_acgpow:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   635
   "has_fpath G p
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   636
  \<Longrightarrow> has_edge (G ^ length (snd p)) (fst p) (prod p) (end_node p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   637
by (cases p)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   638
   (rule path_acgpow_aux[of "snd p" "length (snd p)" _ "fst p", simplified])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   639
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   640
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   641
lemma star_paths:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   642
  "has_edge (star G) a x b =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   643
   (\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   644
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   645
  assume "has_edge (star G) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   646
  then obtain n where pow: "has_edge (G ^ n) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   647
    by (auto simp:in_star)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   648
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   649
  then obtain p where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   650
    "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   651
    by (rule power_induces_path)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   652
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   653
  thus "\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   654
    by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   655
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   656
  assume "\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   657
  then obtain p where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   658
    "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   659
    by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   660
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   661
  hence "has_edge (G ^ length (snd p)) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   662
    by (auto intro:path_acgpow)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   663
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   664
  thus "has_edge (star G) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   665
    by (auto simp:in_star)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   666
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   667
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   668
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   669
lemma plus_paths:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   670
  "has_edge (tcl G) a x b =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   671
   (\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p \<and> 0 < length (snd p))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   672
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   673
  assume "has_edge (tcl G) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   674
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   675
  then obtain n where pow: "has_edge (G ^ n) a x b" and "0 < n"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   676
    by (auto simp:in_tcl)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   677
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   678
  from pow obtain p where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   679
    "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   680
    "n = length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   681
    by (rule power_induces_path)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   682
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   683
  with `0 < n`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   684
  show "\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p \<and> 0 < length (snd p) "
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   685
    by blast
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   686
next
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   687
  assume "\<exists>p. has_fpath G p \<and> a = fst p \<and> b = end_node p \<and> x = prod p
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   688
    \<and> 0 < length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   689
  then obtain p where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   690
    "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   691
    "0 < length (snd p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   692
    by blast
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
  hence "has_edge (G ^ length (snd p)) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   695
    by (auto intro:path_acgpow)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   696
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   697
  with `0 < length (snd p)`
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   698
  show "has_edge (tcl G) a x b"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   699
    by (auto simp:in_tcl)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   700
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   701
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   702
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   703
definition
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   704
  "contract s p = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   705
  (\<lambda>i. (fst (p (s i)), prod (p\<langle>s i,s (Suc i)\<rangle>)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   706
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   707
lemma ipath_contract:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   708
  assumes [simp]: "increasing s"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   709
  assumes ipath: "has_ipath G p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   710
  shows "has_ipath (tcl G) (contract s p)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   711
  unfolding has_ipath_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   712
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   713
  fix i
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   714
  let ?p = "p\<langle>s i,s (Suc i)\<rangle>"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   715
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   716
  from increasing_strict 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   717
	have "fst (p (s (Suc i))) = end_node ?p" by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   718
  moreover
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   719
  from increasing_strict[of s i "Suc i"] have "snd ?p \<noteq> []"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   720
    by (simp add:sub_path_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   721
  moreover
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   722
  from ipath increasing_weak[of s] have "has_fpath G ?p"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   723
    by (rule sub_path_is_path) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   724
  ultimately
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   725
  show "has_edge (tcl G) 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   726
    (fst (contract s p i)) (snd (contract s p i)) (fst (contract s p (Suc i)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   727
    unfolding contract_def plus_paths
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   728
    by (intro exI) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   729
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   730
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   731
lemma prod_unfold:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   732
  "i < j \<Longrightarrow> prod (p\<langle>i,j\<rangle>) 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   733
  = snd (p i) * prod (p\<langle>Suc i, j\<rangle>)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   734
  unfolding prod_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   735
  by (simp add:sub_path_def upt_rec[of "i" j])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   736
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   737
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   738
lemma sub_path_loop:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   739
  assumes "0 < k"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   740
  assumes k: "k = length (snd loop)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   741
  assumes loop: "fst loop = end_node loop"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   742
  shows "(omega loop)\<langle>k * i,k * Suc i\<rangle> = loop" (is "?\<omega> = loop")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   743
proof (rule prod_eqI)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   744
  show "fst ?\<omega> = fst loop"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   745
    by (auto simp:path_loop_def path_nth_def split_def k)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   746
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   747
  show "snd ?\<omega> = snd loop"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   748
  proof (rule nth_equalityI[rule_format])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   749
    show leneq: "length (snd ?\<omega>) = length (snd loop)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   750
      unfolding sub_path_def k by simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   751
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   752
    fix j assume "j < length (snd (?\<omega>))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   753
    with leneq and k have "j < k" by simp
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
    have a: "\<And>i. fst (path_nth loop (Suc i mod k))
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   756
      = snd (snd (path_nth loop (i mod k)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   757
      unfolding k
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   758
      apply (rule path_loop_connect[OF loop])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   759
      using `0 < k` and k
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   760
      apply auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   761
      done
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   762
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   763
    from `j < k` 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   764
    show "snd ?\<omega> ! j = snd loop ! j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   765
      unfolding sub_path_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   766
      apply (simp add:path_loop_def split_def add_ac)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   767
      apply (simp add:a k[symmetric])
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   768
      apply (simp add:path_nth_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   769
      done
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   770
  qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   771
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   772
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   773
end