src/HOL/SizeChange/Size_Change_Termination.thy
author ballarin
Fri, 19 Dec 2008 16:39:23 +0100
changeset 29249 4dc278c8dc59
parent 25314 5eaf3e8b50a4
child 30967 b5d67f83576e
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/Size_Change_Termination.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 "Size-Change Termination"
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 Size_Change_Termination
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
     9
imports Correctness Interpretation Implementation 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    10
uses "sct.ML"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    11
begin
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    12
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    13
subsection {* Simplifier setup *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    14
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    15
text {* This is needed to run the SCT algorithm in the simplifier: *}
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    16
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    17
lemma setbcomp_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    18
  "{x\<in>{}. P x} = {}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    19
  "{x\<in>insert y ys. P x} = (if P y then insert y {x\<in>ys. P x} else {x\<in>ys. P x})"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    20
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    21
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    22
lemma setbcomp_cong:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    23
  "A = B \<Longrightarrow> (\<And>x. P x = Q x) \<Longrightarrow> {x\<in>A. P x} = {x\<in>B. Q x}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    24
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    25
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    26
lemma cartprod_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    27
  "{} \<times> A = {}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    28
  "insert a A \<times> B = Pair a ` B \<union> (A \<times> B)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    29
  by (auto simp:image_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    30
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    31
lemma image_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    32
  "fu ` {} = {}"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    33
  "fu ` insert a A = insert (fu a) (fu ` A)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    34
  by (auto simp:image_def)
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    35
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    36
lemmas union_simps = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    37
  Un_empty_left Un_empty_right Un_insert_left
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    38
  
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    39
lemma subset_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    40
  "{} \<subseteq> B"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    41
  "insert a A \<subseteq> B \<equiv> a \<in> B \<and> A \<subseteq> B"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    42
  by 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
lemma element_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    45
  "x \<in> {} \<equiv> False"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    46
  "x \<in> insert a A \<equiv> x = a \<or> x \<in> A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    47
  by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    48
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    49
lemma set_eq_simp:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    50
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A" by auto
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
lemma ball_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    53
  "\<forall>x\<in>{}. P x \<equiv> True"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    54
  "(\<forall>x\<in>insert a A. P x) \<equiv> P a \<and> (\<forall>x\<in>A. P x)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    55
by auto
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    56
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    57
lemma bex_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    58
  "\<exists>x\<in>{}. P x \<equiv> False"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    59
  "(\<exists>x\<in>insert a A. P x) \<equiv> P a \<or> (\<exists>x\<in>A. P x)"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    60
by auto
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
lemmas set_simps =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    63
  setbcomp_simps
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    64
  cartprod_simps image_simps union_simps subset_simps
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    65
  element_simps set_eq_simp
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    66
  ball_simps bex_simps
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    67
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    68
lemma sedge_simps:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    69
  "\<down> * x = \<down>"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    70
  "\<Down> * x = x"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    71
  by (auto simp:mult_sedge_def)
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
lemmas sctTest_simps =
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    74
  simp_thms
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    75
  if_True
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    76
  if_False
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    77
  nat.inject
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    78
  nat.distinct
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    79
  Pair_eq 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    80
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    81
  grcomp_code 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    82
  edges_match.simps 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    83
  connect_edges.simps 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    84
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    85
  sedge_simps
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    86
  sedge.distinct
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    87
  set_simps
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
  graph_mult_def 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    90
  graph_leq_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    91
  dest_graph.simps
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    92
  graph_plus_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    93
  graph.inject
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    94
  graph_zero_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    95
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    96
  test_SCT_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
    97
  mk_tcl_code
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
  Let_def
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   100
  split_conv
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   101
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   102
lemmas sctTest_congs = 
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   103
  if_weak_cong let_weak_cong setbcomp_cong
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   104
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   105
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   106
lemma SCT_Main:
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   107
  "finite_acg A \<Longrightarrow> test_SCT A \<Longrightarrow> SCT A"
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   108
  using LJA_Theorem4 SCT'_exec
5eaf3e8b50a4 moved stuff about size change termination to its own session
krauss
parents:
diff changeset
   109
  by auto
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