src/HOL/SizeChange/Criterion.thy
author ballarin
Fri, 19 Dec 2008 16:39:23 +0100
changeset 29249 4dc278c8dc59
parent 26513 6f306c8c2c54
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/SCT_Definition.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 {* The Size-Change Principle (Definition) *}
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 Criterion
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     9
imports Graphs Infinite_Set
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 {* 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 sedge =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
    LESS ("\<down>")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
  | LEQ ("\<Down>")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    17
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    18
instantiation sedge :: comm_monoid_mult
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    19
begin
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    21
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    22
  one_sedge_def: "1 = \<Down>"
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    24
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    25
  mult_sedge_def:" a * b = (if a = \<down> then \<down> else b)"
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    26
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    27
instance  proof
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
  fix a b c :: sedge
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
  show "a * b * c = a * (b * c)" by (simp add:mult_sedge_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
  show "1 * a = a" by (simp add:mult_sedge_def one_sedge_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
  show "a * b = b * a" unfolding mult_sedge_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
    by (cases a, simp) (cases b, auto)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    35
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25314
diff changeset
    36
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
lemma sedge_UNIV:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
  "UNIV = { LESS, LEQ }"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
proof (intro equalityI subsetI)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
  fix x show "x \<in> { LESS, LEQ }"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
    by (cases x) auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
qed auto
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
instance sedge :: finite
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
proof
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
  show "finite (UNIV::sedge set)"
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 26441
diff changeset
    47
  by (simp add: sedge_UNIV)
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
qed
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    51
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    52
types 'a scg = "('a, sedge) graph"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
types 'a acg = "('a, 'a scg) graph"
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
subsection {* Size-Change Termination *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    57
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    58
abbreviation (input)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
  "desc P Q == ((\<exists>n.\<forall>i\<ge>n. P i) \<and> (\<exists>\<^sub>\<infinity>i. Q i))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    61
abbreviation (input)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    62
  "dsc G i j \<equiv> has_edge G i LESS j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    63
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    64
abbreviation (input)
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 26441
diff changeset
    65
  "eqp G i j \<equiv> has_edge G i LEQ j"
25314
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    66
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    67
abbreviation
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    68
  eql :: "'a scg \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    69
("_ \<turnstile> _ \<leadsto> _")
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    70
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    71
  "eql G i j \<equiv> has_edge G i LESS j \<or> has_edge G i LEQ j"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    72
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    73
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
abbreviation (input) descat :: "('a, 'a scg) ipath \<Rightarrow> 'a sequence \<Rightarrow> nat \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
  "descat p \<theta> i \<equiv> has_edge (snd (p i)) (\<theta> i) LESS (\<theta> (Suc i))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    77
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
abbreviation (input) eqat :: "('a, 'a scg) ipath \<Rightarrow> 'a sequence \<Rightarrow> nat \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    80
  "eqat p \<theta> i \<equiv> has_edge (snd (p i)) (\<theta> i) LEQ (\<theta> (Suc i))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    81
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    82
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    83
abbreviation (input) eqlat :: "('a, 'a scg) ipath \<Rightarrow> 'a sequence \<Rightarrow> nat \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    84
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    85
  "eqlat p \<theta> i \<equiv> (has_edge (snd (p i)) (\<theta> i) LESS (\<theta> (Suc i))
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    86
                  \<or> has_edge (snd (p i)) (\<theta> i) LEQ (\<theta> (Suc i)))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    87
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    88
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    89
definition is_desc_thread :: "'a sequence \<Rightarrow> ('a, 'a scg) ipath \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    90
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    91
  "is_desc_thread \<theta> p = ((\<exists>n.\<forall>i\<ge>n. eqlat p \<theta> i) \<and> (\<exists>\<^sub>\<infinity>i. descat p \<theta> i))" 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    92
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    93
definition SCT :: "'a acg \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    94
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    95
  "SCT \<A> = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    96
  (\<forall>p. has_ipath \<A> p \<longrightarrow> (\<exists>\<theta>. is_desc_thread \<theta> p))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    97
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    98
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    99
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   100
definition no_bad_graphs :: "'a acg \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   101
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   102
  "no_bad_graphs A = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   103
  (\<forall>n G. has_edge A n G n \<and> G * G = G
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   104
  \<longrightarrow> (\<exists>p. has_edge G p LESS p))"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   105
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   106
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   107
definition SCT' :: "'a acg \<Rightarrow> bool"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   108
where
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   109
  "SCT' A = no_bad_graphs (tcl A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   110
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   111
end