| author | wenzelm | 
| Thu, 06 Dec 2001 00:37:59 +0100 | |
| changeset 12395 | d6913de7655f | 
| parent 12011 | 1a3a7b3cd9bb | 
| child 13915 | 28ccb51bd2f3 | 
| permissions | -rw-r--r-- | 
| 5261 | 1 | (* Title: HOL/Lambda/ListApplication.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1998 TU Muenchen | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 5 | *) | 
| 5261 | 6 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 7 | header {* Application of a term to a list of terms *}
 | 
| 5261 | 8 | |
| 9771 | 9 | theory ListApplication = Lambda: | 
| 10 | ||
| 11 | syntax | |
| 12011 | 12 | "_list_application" :: "dB => dB list => dB" (infixl "\<degree>\<degree>" 150) | 
| 13 | translations | |
| 14 | "t \<degree>\<degree> ts" == "foldl (op \<degree>) t ts" | |
| 11943 | 15 | |
| 12011 | 16 | lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)" | 
| 9771 | 17 | apply (induct_tac ts rule: rev_induct) | 
| 18 | apply auto | |
| 19 | done | |
| 20 | ||
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 21 | lemma Var_eq_apps_conv [rule_format, iff]: | 
| 12011 | 22 | "\<forall>s. (Var m = s \<degree>\<degree> ss) = (Var m = s \<and> ss = [])" | 
| 9771 | 23 | apply (induct_tac ss) | 
| 24 | apply auto | |
| 25 | done | |
| 26 | ||
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 27 | lemma Var_apps_eq_Var_apps_conv [rule_format, iff]: | 
| 12011 | 28 | "\<forall>ss. (Var m \<degree>\<degree> rs = Var n \<degree>\<degree> ss) = (m = n \<and> rs = ss)" | 
| 9771 | 29 | apply (induct_tac rs rule: rev_induct) | 
| 30 | apply simp | |
| 31 | apply blast | |
| 32 | apply (rule allI) | |
| 33 | apply (induct_tac ss rule: rev_induct) | |
| 34 | apply auto | |
| 35 | done | |
| 36 | ||
| 37 | lemma App_eq_foldl_conv: | |
| 12011 | 38 | "(r \<degree> s = t \<degree>\<degree> ts) = | 
| 39 | (if ts = [] then r \<degree> s = t | |
| 40 | else (\<exists>ss. ts = ss @ [s] \<and> r = t \<degree>\<degree> ss))" | |
| 9771 | 41 | apply (rule_tac xs = ts in rev_exhaust) | 
| 42 | apply auto | |
| 43 | done | |
| 44 | ||
| 45 | lemma Abs_eq_apps_conv [iff]: | |
| 12011 | 46 | "(Abs r = s \<degree>\<degree> ss) = (Abs r = s \<and> ss = [])" | 
| 9771 | 47 | apply (induct_tac ss rule: rev_induct) | 
| 48 | apply auto | |
| 49 | done | |
| 50 | ||
| 12011 | 51 | lemma apps_eq_Abs_conv [iff]: "(s \<degree>\<degree> ss = Abs r) = (s = Abs r \<and> ss = [])" | 
| 9771 | 52 | apply (induct_tac ss rule: rev_induct) | 
| 53 | apply auto | |
| 54 | done | |
| 55 | ||
| 56 | lemma Abs_apps_eq_Abs_apps_conv [iff]: | |
| 12011 | 57 | "\<forall>ss. (Abs r \<degree>\<degree> rs = Abs s \<degree>\<degree> ss) = (r = s \<and> rs = ss)" | 
| 9771 | 58 | apply (induct_tac rs rule: rev_induct) | 
| 59 | apply simp | |
| 60 | apply blast | |
| 61 | apply (rule allI) | |
| 62 | apply (induct_tac ss rule: rev_induct) | |
| 63 | apply auto | |
| 64 | done | |
| 65 | ||
| 66 | lemma Abs_App_neq_Var_apps [iff]: | |
| 12011 | 67 | "\<forall>s t. Abs s \<degree> t ~= Var n \<degree>\<degree> ss" | 
| 9771 | 68 | apply (induct_tac ss rule: rev_induct) | 
| 69 | apply auto | |
| 70 | done | |
| 71 | ||
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 72 | lemma Var_apps_neq_Abs_apps [rule_format, iff]: | 
| 12011 | 73 | "\<forall>ts. Var n \<degree>\<degree> ts ~= Abs r \<degree>\<degree> ss" | 
| 9771 | 74 | apply (induct_tac ss rule: rev_induct) | 
| 75 | apply simp | |
| 76 | apply (rule allI) | |
| 77 | apply (induct_tac ts rule: rev_induct) | |
| 78 | apply auto | |
| 79 | done | |
| 5261 | 80 | |
| 9771 | 81 | lemma ex_head_tail: | 
| 12011 | 82 | "\<exists>ts h. t = h \<degree>\<degree> ts \<and> ((\<exists>n. h = Var n) \<or> (\<exists>u. h = Abs u))" | 
| 9771 | 83 | apply (induct_tac t) | 
| 84 | apply (rule_tac x = "[]" in exI) | |
| 85 | apply simp | |
| 86 | apply clarify | |
| 87 | apply (rename_tac ts1 ts2 h1 h2) | |
| 12011 | 88 | apply (rule_tac x = "ts1 @ [h2 \<degree>\<degree> ts2]" in exI) | 
| 9771 | 89 | apply simp | 
| 90 | apply simp | |
| 91 | done | |
| 92 | ||
| 93 | lemma size_apps [simp]: | |
| 12011 | 94 | "size (r \<degree>\<degree> rs) = size r + foldl (op +) 0 (map size rs) + length rs" | 
| 9771 | 95 | apply (induct_tac rs rule: rev_induct) | 
| 96 | apply auto | |
| 97 | done | |
| 98 | ||
| 99 | lemma lem0: "[| (0::nat) < k; m <= n |] ==> m < n + k" | |
| 100 | apply simp | |
| 101 | done | |
| 102 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 103 | |
| 12011 | 104 | text {* \medskip A customized induction schema for @{text "\<degree>\<degree>"}. *}
 | 
| 9771 | 105 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 106 | lemma lem [rule_format (no_asm)]: | 
| 12011 | 107 | "[| !!n ts. \<forall>t \<in> set ts. P t ==> P (Var n \<degree>\<degree> ts); | 
| 108 | !!u ts. [| P u; \<forall>t \<in> set ts. P t |] ==> P (Abs u \<degree>\<degree> ts) | |
| 9771 | 109 | |] ==> \<forall>t. size t = n --> P t" | 
| 110 | proof - | |
| 11549 | 111 | case rule_context | 
| 9771 | 112 | show ?thesis | 
| 9870 | 113 | apply (induct_tac n rule: nat_less_induct) | 
| 9771 | 114 | apply (rule allI) | 
| 115 | apply (cut_tac t = t in ex_head_tail) | |
| 116 | apply clarify | |
| 117 | apply (erule disjE) | |
| 118 | apply clarify | |
| 119 | apply (rule prems) | |
| 120 | apply clarify | |
| 121 | apply (erule allE, erule impE) | |
| 122 | prefer 2 | |
| 123 | apply (erule allE, erule mp, rule refl) | |
| 124 | apply simp | |
| 125 | apply (rule lem0) | |
| 126 | apply force | |
| 127 | apply (rule elem_le_sum) | |
| 128 | apply force | |
| 129 | apply clarify | |
| 130 | apply (rule prems) | |
| 131 | apply (erule allE, erule impE) | |
| 132 | prefer 2 | |
| 133 | apply (erule allE, erule mp, rule refl) | |
| 134 | apply simp | |
| 135 | apply clarify | |
| 136 | apply (erule allE, erule impE) | |
| 137 | prefer 2 | |
| 138 | apply (erule allE, erule mp, rule refl) | |
| 139 | apply simp | |
| 140 | apply (rule le_imp_less_Suc) | |
| 141 | apply (rule trans_le_add1) | |
| 142 | apply (rule trans_le_add2) | |
| 143 | apply (rule elem_le_sum) | |
| 144 | apply force | |
| 145 | done | |
| 146 | qed | |
| 147 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 148 | theorem Apps_dB_induct: | 
| 12011 | 149 | "[| !!n ts. \<forall>t \<in> set ts. P t ==> P (Var n \<degree>\<degree> ts); | 
| 150 | !!u ts. [| P u; \<forall>t \<in> set ts. P t |] ==> P (Abs u \<degree>\<degree> ts) | |
| 9771 | 151 | |] ==> P t" | 
| 152 | proof - | |
| 11549 | 153 | case rule_context | 
| 9771 | 154 | show ?thesis | 
| 155 | apply (rule_tac t = t in lem) | |
| 156 | prefer 3 | |
| 157 | apply (rule refl) | |
| 158 | apply (assumption | rule prems)+ | |
| 159 | done | |
| 160 | qed | |
| 5261 | 161 | |
| 9870 | 162 | end |