Up to index of Isabelle/HOL/SizeChange
theory Size_Change_Termination(* Title: HOL/Library/Size_Change_Termination.thy ID: $Id: Size_Change_Termination.thy,v 1.1 2007/11/06 16:44:54 krauss Exp $ Author: Alexander Krauss, TU Muenchen *) header "Size-Change Termination" theory Size_Change_Termination imports Correctness Interpretation Implementation uses "sct.ML" begin subsection {* Simplifier setup *} text {* This is needed to run the SCT algorithm in the simplifier: *} lemma setbcomp_simps: "{x∈{}. P x} = {}" "{x∈insert y ys. P x} = (if P y then insert y {x∈ys. P x} else {x∈ys. P x})" by auto lemma setbcomp_cong: "A = B ==> (!!x. P x = Q x) ==> {x∈A. P x} = {x∈B. Q x}" by auto lemma cartprod_simps: "{} × A = {}" "insert a A × B = Pair a ` B ∪ (A × B)" by (auto simp:image_def) lemma image_simps: "fu ` {} = {}" "fu ` insert a A = insert (fu a) (fu ` A)" by (auto simp:image_def) lemmas union_simps = Un_empty_left Un_empty_right Un_insert_left lemma subset_simps: "{} ⊆ B" "insert a A ⊆ B ≡ a ∈ B ∧ A ⊆ B" by auto lemma element_simps: "x ∈ {} ≡ False" "x ∈ insert a A ≡ x = a ∨ x ∈ A" by auto lemma set_eq_simp: "A = B <-> A ⊆ B ∧ B ⊆ A" by auto lemma ball_simps: "∀x∈{}. P x ≡ True" "(∀x∈insert a A. P x) ≡ P a ∧ (∀x∈A. P x)" by auto lemma bex_simps: "∃x∈{}. P x ≡ False" "(∃x∈insert a A. P x) ≡ P a ∨ (∃x∈A. P x)" by auto lemmas set_simps = setbcomp_simps cartprod_simps image_simps union_simps subset_simps element_simps set_eq_simp ball_simps bex_simps lemma sedge_simps: "\<down> * x = \<down>" "\<Down> * x = x" by (auto simp:mult_sedge_def) lemmas sctTest_simps = simp_thms if_True if_False nat.inject nat.distinct Pair_eq grcomp_code edges_match.simps connect_edges.simps sedge_simps sedge.distinct set_simps graph_mult_def graph_leq_def dest_graph.simps graph_plus_def graph.inject graph_zero_def test_SCT_def mk_tcl_code Let_def split_conv lemmas sctTest_congs = if_weak_cong let_weak_cong setbcomp_cong lemma SCT_Main: "finite_acg A ==> test_SCT A ==> SCT A" using LJA_Theorem4 SCT'_exec by auto end
lemma setbcomp_simps:
{x : {}. P x} = {}
{x : insert y ys. P x} = (if P y then insert y {x : ys. P x} else {x : ys. P x})
lemma setbcomp_cong:
[| A = B; !!x. P x = Q x |] ==> {x : A. P x} = {x : B. Q x}
lemma cartprod_simps:
{} × A = {}
insert a A × B = Pair a ` B ∪ A × B
lemma image_simps:
fu ` {} = {}
fu ` insert a A = insert (fu a) (fu ` A)
lemma union_simps:
{} ∪ B = B
A ∪ {} = A
insert a B ∪ C = insert a (B ∪ C)
lemma subset_simps:
{} ⊆ B
insert a A ⊆ B == a ∈ B ∧ A ⊆ B
lemma element_simps:
x ∈ {} == False
x ∈ insert a A == x = a ∨ x ∈ A
lemma set_eq_simp:
(A = B) = (A ⊆ B ∧ B ⊆ A)
lemma ball_simps:
∀x∈{}. P x == True
∀x∈insert a A. P x == P a ∧ (∀x∈A. P x)
lemma bex_simps:
∃x∈{}. P x == False
∃x∈insert a A. P x == P a ∨ (∃x∈A. P x)
lemma set_simps:
{x : {}. P x} = {}
{x : insert y ys. P x} = (if P y then insert y {x : ys. P x} else {x : ys. P x})
{} × A = {}
insert a A × B = Pair a ` B ∪ A × B
fu ` {} = {}
fu ` insert a A = insert (fu a) (fu ` A)
{} ∪ B = B
A ∪ {} = A
insert a B ∪ C = insert a (B ∪ C)
{} ⊆ B
insert a A ⊆ B == a ∈ B ∧ A ⊆ B
x ∈ {} == False
x ∈ insert a A == x = a ∨ x ∈ A
(A = B) = (A ⊆ B ∧ B ⊆ A)
∀x∈{}. P x == True
∀x∈insert a A. P x == P a ∧ (∀x∈A. P x)
∃x∈{}. P x == False
∃x∈insert a A. P x == P a ∨ (∃x∈A. P x)
lemma sedge_simps:
\<down> * x = \<down>
\<Down> * x = x
lemma sctTest_simps:
(¬ ¬ P) = P
((¬ P) = (¬ Q)) = (P = Q)
(P ≠ Q) = (P = (¬ Q))
(P ∨ ¬ P) = True
(¬ P ∨ P) = True
(x = x) = True
(¬ True) = False
(¬ False) = True
(¬ P) ≠ P
P ≠ (¬ P)
(True = P) = P
(P = True) = P
(False = P) = (¬ P)
(P = False) = (¬ P)
(True --> P) = P
(False --> P) = True
(P --> True) = True
(P --> P) = True
(P --> False) = (¬ P)
(P --> ¬ P) = (¬ P)
(P ∧ True) = P
(True ∧ P) = P
(P ∧ False) = False
(False ∧ P) = False
(P ∧ P) = P
(P ∧ P ∧ Q) = (P ∧ Q)
(P ∧ ¬ P) = False
(¬ P ∧ P) = False
(P ∨ True) = True
(True ∨ P) = True
(P ∨ False) = P
(False ∨ P) = P
(P ∨ P) = P
(P ∨ P ∨ Q) = (P ∨ Q)
(∀x. P) = P
(∃x. P) = P
∃x. x = t
∃x. t = x
(∃x. x = t ∧ P x) = P t
(∃x. t = x ∧ P x) = P t
(∀x. x = t --> P x) = P t
(∀x. t = x --> P x) = P t
(if True then x else y) = x
(if False then x else y) = y
(Suc m = Suc n) = (m = n)
Suc m ≠ 0
0 ≠ Suc m
((a, b) = (a', b')) = (a = a' ∧ b = b')
grcomp (Graph G) (Graph H) = Graph (connect_edges ` {x : G × H. edges_match x})
edges_match ((n, e, m), n', e', m') = (m = n')
connect_edges ((n, e, m), n', e', m') = (n, e * e', m')
\<down> * x = \<down>
\<Down> * x = x
\<down> ≠ \<Down>
\<Down> ≠ \<down>
{x : {}. P x} = {}
{x : insert y ys. P x} = (if P y then insert y {x : ys. P x} else {x : ys. P x})
{} × A = {}
insert a A × B = Pair a ` B ∪ A × B
fu ` {} = {}
fu ` insert a A = insert (fu a) (fu ` A)
{} ∪ B = B
A ∪ {} = A
insert a B ∪ C = insert a (B ∪ C)
{} ⊆ B
insert a A ⊆ B == a ∈ B ∧ A ⊆ B
x ∈ {} == False
x ∈ insert a A == x = a ∨ x ∈ A
(A = B) = (A ⊆ B ∧ B ⊆ A)
∀x∈{}. P x == True
∀x∈insert a A. P x == P a ∧ (∀x∈A. P x)
∃x∈{}. P x == False
∃x∈insert a A. P x == P a ∨ (∃x∈A. P x)
G * H == grcomp G H
G ≤ H == dest_graph G ⊆ dest_graph H
dest_graph (Graph G) = G
G + H == Graph (dest_graph G ∪ dest_graph H)
(Graph set = Graph set') = (set = set')
0 == Graph {}
test_SCT \<A> =
(let \<T> = mk_tcl \<A> \<A>
in ∀(n, G, m)∈dest_graph \<T>.
n ≠ m ∨ G * G ≠ G ∨ (∃(p, e, q)∈dest_graph G. p = q ∧ e = \<down>))
mk_tcl A X = (let XA = X * A in if XA ≤ X then X else mk_tcl A (X + XA))
Let s f == f s
split c (a, b) = c a b
lemma sctTest_congs:
b = c ==> (if b then x else y) = (if c then x else y)
a = b ==> Let a t = Let b t
[| A = B; !!x. P x = Q x |] ==> {x : A. P x} = {x : B. Q x}
lemma SCT_Main:
[| finite_acg A; test_SCT A |] ==> SCT A