| author | haftmann | 
| Tue, 26 May 2009 17:29:33 +0200 | |
| changeset 31261 | 900ebbc35e30 | 
| parent 30967 | b5d67f83576e | 
| permissions | -rw-r--r-- | 
| 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 | Author: Alexander Krauss, TU Muenchen | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 3 | *) | 
| 
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 | header "Size-Change Termination" | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 6 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 7 | theory Size_Change_Termination | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 8 | imports Correctness Interpretation Implementation | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 9 | uses "sct.ML" | 
| 
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 {* Simplifier setup *}
 | 
| 
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 | 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 | 15 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 16 | lemma setbcomp_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 17 |   "{x\<in>{}. P x} = {}"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 18 |   "{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 | 19 | by auto | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 20 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 21 | lemma setbcomp_cong: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 22 |   "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 | 23 | by auto | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 24 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 25 | lemma cartprod_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 26 |   "{} \<times> A = {}"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 27 | "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 | 28 | by (auto simp:image_def) | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 29 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 30 | lemma image_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 31 |   "fu ` {} = {}"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 32 | "fu ` insert a A = insert (fu a) (fu ` A)" | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 33 | by (auto simp:image_def) | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 34 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 35 | lemmas union_simps = | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 36 | Un_empty_left Un_empty_right Un_insert_left | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 37 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 38 | lemma subset_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 39 |   "{} \<subseteq> B"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 40 | "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 | 41 | by auto | 
| 
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 | lemma element_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 44 |   "x \<in> {} \<equiv> False"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 45 | "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 | 46 | by auto | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 47 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 48 | lemma set_eq_simp: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 49 | "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 | 50 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 51 | lemma ball_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 52 |   "\<forall>x\<in>{}. P x \<equiv> True"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 53 | "(\<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 | 54 | by auto | 
| 
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 bex_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 57 |   "\<exists>x\<in>{}. P x \<equiv> False"
 | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 58 | "(\<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 | 59 | by auto | 
| 
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 | lemmas set_simps = | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 62 | setbcomp_simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 63 | cartprod_simps image_simps union_simps subset_simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 64 | element_simps set_eq_simp | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 65 | ball_simps bex_simps | 
| 
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 | lemma sedge_simps: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 68 | "\<down> * x = \<down>" | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 69 | "\<Down> * x = x" | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 70 | by (auto simp:mult_sedge_def) | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 71 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 72 | lemmas sctTest_simps = | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 73 | simp_thms | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 74 | if_True | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 75 | if_False | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 76 | nat.inject | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 77 | nat.distinct | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 78 | Pair_eq | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 79 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 80 | grcomp_code | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 81 | edges_match.simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 82 | connect_edges.simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 83 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 84 | sedge_simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 85 | sedge.distinct | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 86 | set_simps | 
| 
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 | graph_mult_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 89 | graph_leq_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 90 | dest_graph.simps | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 91 | graph_plus_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 92 | graph.inject | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 93 | graph_zero_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 94 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 95 | test_SCT_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 96 | mk_tcl_code | 
| 
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 | Let_def | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 99 | split_conv | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 100 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 101 | lemmas sctTest_congs = | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 102 | if_weak_cong let_weak_cong setbcomp_cong | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 103 | |
| 
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 | lemma SCT_Main: | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 106 | "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 | 107 | using LJA_Theorem4 SCT'_exec | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 108 | by auto | 
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 109 | |
| 
5eaf3e8b50a4
moved stuff about size change termination to its own session
 krauss parents: diff
changeset | 110 | end |