Theory Criterion

Up to index of Isabelle/HOL/SizeChange

theory Criterion
imports Graphs Infinite_Set
begin

(*  Title:      HOL/Library/SCT_Definition.thy
    ID:         $Id: Criterion.thy,v 1.4 2008/04/02 13:58:34 haftmann Exp $
    Author:     Alexander Krauss, TU Muenchen
*)

header {* The Size-Change Principle (Definition) *}

theory Criterion
imports Graphs Infinite_Set
begin

subsection {* Size-Change Graphs *}

datatype sedge =
    LESS ("\<down>")
  | LEQ ("\<Down>")

instantiation sedge :: comm_monoid_mult
begin

definition
  one_sedge_def: "1 = \<Down>"

definition
  mult_sedge_def:" a * b = (if a = \<down> then \<down> else b)"

instance  proof
  fix a b c :: sedge
  show "a * b * c = a * (b * c)" by (simp add:mult_sedge_def)
  show "1 * a = a" by (simp add:mult_sedge_def one_sedge_def)
  show "a * b = b * a" unfolding mult_sedge_def
    by (cases a, simp) (cases b, auto)
qed

end

lemma sedge_UNIV:
  "UNIV = { LESS, LEQ }"
proof (intro equalityI subsetI)
  fix x show "x ∈ { LESS, LEQ }"
    by (cases x) auto
qed auto

instance sedge :: finite
proof
  show "finite (UNIV::sedge set)"
  by (simp add: sedge_UNIV)
qed



types 'a scg = "('a, sedge) graph"
types 'a acg = "('a, 'a scg) graph"


subsection {* Size-Change Termination *}

abbreviation (input)
  "desc P Q == ((∃n.∀i≥n. P i) ∧ (∃i. Q i))"

abbreviation (input)
  "dsc G i j ≡ has_edge G i LESS j"

abbreviation (input)
  "eqp G i j ≡ has_edge G i LEQ j"

abbreviation
  eql :: "'a scg => 'a => 'a => bool"
("_ \<turnstile> _ \<leadsto> _")
where
  "eql G i j ≡ has_edge G i LESS j ∨ has_edge G i LEQ j"


abbreviation (input) descat :: "('a, 'a scg) ipath => 'a sequence => nat => bool"
where
  "descat p ϑ i ≡ has_edge (snd (p i)) (ϑ i) LESS (ϑ (Suc i))"

abbreviation (input) eqat :: "('a, 'a scg) ipath => 'a sequence => nat => bool"
where
  "eqat p ϑ i ≡ has_edge (snd (p i)) (ϑ i) LEQ (ϑ (Suc i))"


abbreviation (input) eqlat :: "('a, 'a scg) ipath => 'a sequence => nat => bool"
where
  "eqlat p ϑ i ≡ (has_edge (snd (p i)) (ϑ i) LESS (ϑ (Suc i))
                  ∨ has_edge (snd (p i)) (ϑ i) LEQ (ϑ (Suc i)))"


definition is_desc_thread :: "'a sequence => ('a, 'a scg) ipath => bool"
where
  "is_desc_thread ϑ p = ((∃n.∀i≥n. eqlat p ϑ i) ∧ (∃i. descat p ϑ i))" 

definition SCT :: "'a acg => bool"
where
  "SCT \<A> = 
  (∀p. has_ipath \<A> p --> (∃ϑ. is_desc_thread ϑ p))"



definition no_bad_graphs :: "'a acg => bool"
where
  "no_bad_graphs A = 
  (∀n G. has_edge A n G n ∧ G * G = G
  --> (∃p. has_edge G p LESS p))"


definition SCT' :: "'a acg => bool"
where
  "SCT' A = no_bad_graphs (tcl A)"

end

Size-Change Graphs

lemma sedge_UNIV:

  UNIV = {\<down>, \<Down>}

Size-Change Termination