| author | wenzelm | 
| Wed, 27 Mar 2024 13:23:15 +0100 | |
| changeset 80026 | a03a7d4b82f8 | 
| parent 75287 | 7add2d5322a7 | 
| child 80067 | c40bdfc84640 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Induct/QuoNestedDataType.thy | 
| 15172 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 2004 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 60530 | 6 | section\<open>Quotienting a Free Algebra Involving Nested Recursion\<close> | 
| 15172 | 7 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 8 | text \<open>This is the development promised in Lawrence Paulson's paper ``Defining functions on equivalence classes'' | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 9 | \emph{ACM Transactions on Computational Logic} \textbf{7}:40 (2006), 658--675,
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 10 | illustrating bare-bones quotient constructions. Any comparison using lifting and transfer | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 11 | should be done in a separate theory.\<close> | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 12 | |
| 16417 | 13 | theory QuoNestedDataType imports Main begin | 
| 15172 | 14 | |
| 60530 | 15 | subsection\<open>Defining the Free Algebra\<close> | 
| 15172 | 16 | |
| 60530 | 17 | text\<open>Messages with encryption and decryption as free constructors.\<close> | 
| 58310 | 18 | datatype | 
| 15172 | 19 | freeExp = VAR nat | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30198diff
changeset | 20 | | PLUS freeExp freeExp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30198diff
changeset | 21 | | FNCALL nat "freeExp list" | 
| 15172 | 22 | |
| 58305 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 blanchet parents: 
55417diff
changeset | 23 | datatype_compat freeExp | 
| 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 blanchet parents: 
55417diff
changeset | 24 | |
| 60530 | 25 | text\<open>The equivalence relation, which makes PLUS associative.\<close> | 
| 15172 | 26 | |
| 60530 | 27 | text\<open>The first rule is the desired equation. The next three rules | 
| 15172 | 28 | make the equations applicable to subterms. The last two rules are symmetry | 
| 60530 | 29 | and transitivity.\<close> | 
| 23746 | 30 | inductive_set | 
| 31 | exprel :: "(freeExp * freeExp) set" | |
| 32 | and exp_rel :: "[freeExp, freeExp] => bool" (infixl "\<sim>" 50) | |
| 33 | where | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 34 | "X \<sim> Y \<equiv> (X,Y) \<in> exprel" | 
| 23746 | 35 | | ASSOC: "PLUS X (PLUS Y Z) \<sim> PLUS (PLUS X Y) Z" | 
| 36 | | VAR: "VAR N \<sim> VAR N" | |
| 37 | | PLUS: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> PLUS X Y \<sim> PLUS X' Y'" | |
| 38 | | FNCALL: "(Xs,Xs') \<in> listrel exprel \<Longrightarrow> FNCALL F Xs \<sim> FNCALL F Xs'" | |
| 39 | | SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X" | |
| 40 | | TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z" | |
| 15172 | 41 | monos listrel_mono | 
| 42 | ||
| 43 | ||
| 60530 | 44 | text\<open>Proving that it is an equivalence relation\<close> | 
| 15172 | 45 | |
| 18460 | 46 | lemma exprel_refl: "X \<sim> X" | 
| 47 | and list_exprel_refl: "(Xs,Xs) \<in> listrel(exprel)" | |
| 58305 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 blanchet parents: 
55417diff
changeset | 48 | by (induct X and Xs rule: compat_freeExp.induct compat_freeExp_list.induct) | 
| 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 blanchet parents: 
55417diff
changeset | 49 | (blast intro: exprel.intros listrel.intros)+ | 
| 15172 | 50 | |
| 51 | theorem equiv_exprel: "equiv UNIV exprel" | |
| 18460 | 52 | proof - | 
| 30198 | 53 | have "refl exprel" by (simp add: refl_on_def exprel_refl) | 
| 18460 | 54 | moreover have "sym exprel" by (simp add: sym_def, blast intro: exprel.SYM) | 
| 55 | moreover have "trans exprel" by (simp add: trans_def, blast intro: exprel.TRANS) | |
| 56 | ultimately show ?thesis by (simp add: equiv_def) | |
| 15172 | 57 | qed | 
| 58 | ||
| 59 | theorem equiv_list_exprel: "equiv UNIV (listrel exprel)" | |
| 18460 | 60 | using equiv_listrel [OF equiv_exprel] by simp | 
| 15172 | 61 | |
| 62 | lemma FNCALL_Cons: | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 63 | "\<lbrakk>X \<sim> X'; (Xs,Xs') \<in> listrel(exprel)\<rbrakk> \<Longrightarrow> FNCALL F (X#Xs) \<sim> FNCALL F (X'#Xs')" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 64 | by (blast intro: exprel.intros listrel.intros) | 
| 15172 | 65 | |
| 66 | ||
| 60530 | 67 | subsection\<open>Some Functions on the Free Algebra\<close> | 
| 15172 | 68 | |
| 60530 | 69 | subsubsection\<open>The Set of Variables\<close> | 
| 15172 | 70 | |
| 60530 | 71 | text\<open>A function to return the set of variables present in a message. It will | 
| 59478 | 72 | be lifted to the initial algebra, to serve as an example of that process. | 
| 15172 | 73 | Note that the "free" refers to the free datatype rather than to the concept | 
| 60530 | 74 | of a free variable.\<close> | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 75 | primrec freevars :: "freeExp \<Rightarrow> nat set" and freevars_list :: "freeExp list \<Rightarrow> nat set" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 76 | where | 
| 39246 | 77 |   "freevars (VAR N) = {N}"
 | 
| 78 | | "freevars (PLUS X Y) = freevars X \<union> freevars Y" | |
| 79 | | "freevars (FNCALL F Xs) = freevars_list Xs" | |
| 15172 | 80 | |
| 39246 | 81 | | "freevars_list [] = {}"
 | 
| 82 | | "freevars_list (X # Xs) = freevars X \<union> freevars_list Xs" | |
| 15172 | 83 | |
| 60530 | 84 | text\<open>This theorem lets us prove that the vars function respects the | 
| 15172 | 85 | equivalence relation. It also helps us prove that Variable | 
| 60530 | 86 | (the abstract constructor) is injective\<close> | 
| 15172 | 87 | theorem exprel_imp_eq_freevars: "U \<sim> V \<Longrightarrow> freevars U = freevars V" | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 88 | proof (induct set: exprel) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 89 | case (FNCALL Xs Xs' F) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 90 | then show ?case | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 91 | by (induct rule: listrel.induct) auto | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 92 | qed (simp_all add: Un_assoc) | 
| 15172 | 93 | |
| 94 | ||
| 60530 | 95 | subsubsection\<open>Functions for Freeness\<close> | 
| 15172 | 96 | |
| 60530 | 97 | text\<open>A discriminator function to distinguish vars, sums and function calls\<close> | 
| 39246 | 98 | primrec freediscrim :: "freeExp \<Rightarrow> int" where | 
| 99 | "freediscrim (VAR N) = 0" | |
| 100 | | "freediscrim (PLUS X Y) = 1" | |
| 101 | | "freediscrim (FNCALL F Xs) = 2" | |
| 15172 | 102 | |
| 103 | theorem exprel_imp_eq_freediscrim: | |
| 104 | "U \<sim> V \<Longrightarrow> freediscrim U = freediscrim V" | |
| 18460 | 105 | by (induct set: exprel) auto | 
| 15172 | 106 | |
| 107 | ||
| 60530 | 108 | text\<open>This function, which returns the function name, is used to | 
| 109 | prove part of the injectivity property for FnCall.\<close> | |
| 39246 | 110 | primrec freefun :: "freeExp \<Rightarrow> nat" where | 
| 111 | "freefun (VAR N) = 0" | |
| 112 | | "freefun (PLUS X Y) = 0" | |
| 113 | | "freefun (FNCALL F Xs) = F" | |
| 15172 | 114 | |
| 115 | theorem exprel_imp_eq_freefun: | |
| 116 | "U \<sim> V \<Longrightarrow> freefun U = freefun V" | |
| 23746 | 117 | by (induct set: exprel) (simp_all add: listrel.intros) | 
| 15172 | 118 | |
| 119 | ||
| 60530 | 120 | text\<open>This function, which returns the list of function arguments, is used to | 
| 121 | prove part of the injectivity property for FnCall.\<close> | |
| 39246 | 122 | primrec freeargs :: "freeExp \<Rightarrow> freeExp list" where | 
| 123 | "freeargs (VAR N) = []" | |
| 124 | | "freeargs (PLUS X Y) = []" | |
| 125 | | "freeargs (FNCALL F Xs) = Xs" | |
| 15172 | 126 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 127 | |
| 15172 | 128 | theorem exprel_imp_eqv_freeargs: | 
| 40822 | 129 | assumes "U \<sim> V" | 
| 130 | shows "(freeargs U, freeargs V) \<in> listrel exprel" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 131 | using assms | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 132 | proof induction | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 133 | case (FNCALL Xs Xs' F) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 134 | then show ?case | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 135 | by (simp add: listrel_iff_nth) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 136 | next | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 137 | case (SYM X Y) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 138 | then show ?case | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 139 | by (meson equivE equiv_list_exprel symD) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 140 | next | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 141 | case (TRANS X Y Z) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 142 | then show ?case | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 143 | by (meson equivE equiv_list_exprel transD) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 144 | qed (use listrel.simps in auto) | 
| 15172 | 145 | |
| 146 | ||
| 60530 | 147 | subsection\<open>The Initial Algebra: A Quotiented Message Type\<close> | 
| 15172 | 148 | |
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
41959diff
changeset | 149 | definition "Exp = UNIV//exprel" | 
| 15172 | 150 | |
| 49834 | 151 | typedef exp = Exp | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
41959diff
changeset | 152 | morphisms Rep_Exp Abs_Exp | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
41959diff
changeset | 153 | unfolding Exp_def by (auto simp add: quotient_def) | 
| 15172 | 154 | |
| 60530 | 155 | text\<open>The abstract message constructors\<close> | 
| 15172 | 156 | |
| 19736 | 157 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 158 | Var :: "nat \<Rightarrow> exp" where | 
| 19736 | 159 |   "Var N = Abs_Exp(exprel``{VAR N})"
 | 
| 15172 | 160 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 161 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 162 | Plus :: "[exp,exp] \<Rightarrow> exp" where | 
| 19736 | 163 | "Plus X Y = | 
| 15172 | 164 |        Abs_Exp (\<Union>U \<in> Rep_Exp X. \<Union>V \<in> Rep_Exp Y. exprel``{PLUS U V})"
 | 
| 165 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 166 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 167 | FnCall :: "[nat, exp list] \<Rightarrow> exp" where | 
| 19736 | 168 | "FnCall F Xs = | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 169 |        Abs_Exp (\<Union>Us \<in> listset (map Rep_Exp Xs). exprel``{FNCALL F Us})"
 | 
| 15172 | 170 | |
| 171 | ||
| 69597 | 172 | text\<open>Reduces equality of equivalence classes to the \<^term>\<open>exprel\<close> relation: | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 173 |   \<^term>\<open>(exprel``{x} = exprel``{y}) = ((x,y) \<in> exprel)\<close>\<close>
 | 
| 15172 | 174 | lemmas equiv_exprel_iff = eq_equiv_class_iff [OF equiv_exprel UNIV_I UNIV_I] | 
| 175 | ||
| 176 | declare equiv_exprel_iff [simp] | |
| 177 | ||
| 178 | ||
| 60530 | 179 | text\<open>All equivalence classes belong to set of representatives\<close> | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 180 | lemma exprel_in_Exp [simp]: "exprel``{U} \<in> Exp"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 181 | by (simp add: Exp_def quotientI) | 
| 15172 | 182 | |
| 183 | lemma inj_on_Abs_Exp: "inj_on Abs_Exp Exp" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 184 | by (meson Abs_Exp_inject inj_onI) | 
| 15172 | 185 | |
| 60530 | 186 | text\<open>Reduces equality on abstractions to equality on representatives\<close> | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
60530diff
changeset | 187 | declare inj_on_Abs_Exp [THEN inj_on_eq_iff, simp] | 
| 15172 | 188 | |
| 189 | declare Abs_Exp_inverse [simp] | |
| 190 | ||
| 191 | ||
| 60530 | 192 | text\<open>Case analysis on the representation of a exp as an equivalence class.\<close> | 
| 15172 | 193 | lemma eq_Abs_Exp [case_names Abs_Exp, cases type: exp]: | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 194 |      "(\<And>U. z = Abs_Exp (exprel``{U}) \<Longrightarrow> P) \<Longrightarrow> P"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 195 | by (metis Abs_Exp_cases Exp_def quotientE) | 
| 15172 | 196 | |
| 197 | ||
| 60530 | 198 | subsection\<open>Every list of abstract expressions can be expressed in terms of a | 
| 199 | list of concrete expressions\<close> | |
| 15172 | 200 | |
| 19736 | 201 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 202 | Abs_ExpList :: "freeExp list => exp list" where | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 203 |   "Abs_ExpList Xs \<equiv> map (\<lambda>U. Abs_Exp(exprel``{U})) Xs"
 | 
| 15172 | 204 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 205 | lemma Abs_ExpList_Nil [simp]: "Abs_ExpList [] = []" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 206 | by (simp add: Abs_ExpList_def) | 
| 15172 | 207 | |
| 208 | lemma Abs_ExpList_Cons [simp]: | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 209 |   "Abs_ExpList (X#Xs) = Abs_Exp (exprel``{X}) # Abs_ExpList Xs"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 210 | by (simp add: Abs_ExpList_def) | 
| 15172 | 211 | |
| 212 | lemma ExpList_rep: "\<exists>Us. z = Abs_ExpList Us" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 213 | by (smt (verit, del_insts) Abs_ExpList_def eq_Abs_Exp ex_map_conv) | 
| 15172 | 214 | |
| 215 | ||
| 60530 | 216 | subsubsection\<open>Characteristic Equations for the Abstract Constructors\<close> | 
| 15172 | 217 | |
| 218 | lemma Plus: "Plus (Abs_Exp(exprel``{U})) (Abs_Exp(exprel``{V})) = 
 | |
| 219 |              Abs_Exp (exprel``{PLUS U V})"
 | |
| 220 | proof - | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 221 |   have "(\<lambda>U V. exprel``{PLUS U V}) respects2 exprel"
 | 
| 40822 | 222 | by (auto simp add: congruent2_def exprel.PLUS) | 
| 15172 | 223 | thus ?thesis | 
| 224 | by (simp add: Plus_def UN_equiv_class2 [OF equiv_exprel equiv_exprel]) | |
| 225 | qed | |
| 226 | ||
| 60530 | 227 | text\<open>It is not clear what to do with FnCall: it's argument is an abstraction | 
| 69597 | 228 | of an \<^typ>\<open>exp list\<close>. Is it just Nil or Cons? What seems to work best is to | 
| 229 | regard an \<^typ>\<open>exp list\<close> as a \<^term>\<open>listrel exprel\<close> equivalence class\<close> | |
| 15172 | 230 | |
| 60530 | 231 | text\<open>This theorem is easily proved but never used. There's no obvious way | 
| 63167 | 232 | even to state the analogous result, \<open>FnCall_Cons\<close>.\<close> | 
| 15172 | 233 | lemma FnCall_Nil: "FnCall F [] = Abs_Exp (exprel``{FNCALL F []})"
 | 
| 234 | by (simp add: FnCall_def) | |
| 235 | ||
| 236 | lemma FnCall_respects: | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 237 |      "(\<lambda>Us. exprel``{FNCALL F Us}) respects (listrel exprel)"
 | 
| 40822 | 238 | by (auto simp add: congruent_def exprel.FNCALL) | 
| 15172 | 239 | |
| 240 | lemma FnCall_sing: | |
| 241 |      "FnCall F [Abs_Exp(exprel``{U})] = Abs_Exp (exprel``{FNCALL F [U]})"
 | |
| 242 | proof - | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 243 |   have "(\<lambda>U. exprel``{FNCALL F [U]}) respects exprel"
 | 
| 40822 | 244 | by (auto simp add: congruent_def FNCALL_Cons listrel.intros) | 
| 15172 | 245 | thus ?thesis | 
| 246 | by (simp add: FnCall_def UN_equiv_class [OF equiv_exprel]) | |
| 247 | qed | |
| 248 | ||
| 249 | lemma listset_Rep_Exp_Abs_Exp: | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 250 |      "listset (map Rep_Exp (Abs_ExpList Us)) = listrel exprel``{Us}"
 | 
| 18460 | 251 | by (induct Us) (simp_all add: listrel_Cons Abs_ExpList_def) | 
| 15172 | 252 | |
| 253 | lemma FnCall: | |
| 254 |      "FnCall F (Abs_ExpList Us) = Abs_Exp (exprel``{FNCALL F Us})"
 | |
| 255 | proof - | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 256 |   have "(\<lambda>Us. exprel``{FNCALL F Us}) respects (listrel exprel)"
 | 
| 40822 | 257 | by (auto simp add: congruent_def exprel.FNCALL) | 
| 15172 | 258 | thus ?thesis | 
| 259 | by (simp add: FnCall_def UN_equiv_class [OF equiv_list_exprel] | |
| 260 | listset_Rep_Exp_Abs_Exp) | |
| 261 | qed | |
| 262 | ||
| 263 | ||
| 60530 | 264 | text\<open>Establishing this equation is the point of the whole exercise\<close> | 
| 15172 | 265 | theorem Plus_assoc: "Plus X (Plus Y Z) = Plus (Plus X Y) Z" | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 266 | by (cases X, cases Y, cases Z, simp add: Plus exprel.ASSOC) | 
| 15172 | 267 | |
| 268 | ||
| 269 | ||
| 60530 | 270 | subsection\<open>The Abstract Function to Return the Set of Variables\<close> | 
| 15172 | 271 | |
| 19736 | 272 | definition | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 273 | vars :: "exp \<Rightarrow> nat set" where "vars X \<equiv> (\<Union>U \<in> Rep_Exp X. freevars U)" | 
| 15172 | 274 | |
| 275 | lemma vars_respects: "freevars respects exprel" | |
| 40822 | 276 | by (auto simp add: congruent_def exprel_imp_eq_freevars) | 
| 15172 | 277 | |
| 69597 | 278 | text\<open>The extension of the function \<^term>\<open>vars\<close> to lists\<close> | 
| 39246 | 279 | primrec vars_list :: "exp list \<Rightarrow> nat set" where | 
| 280 |   "vars_list []    = {}"
 | |
| 281 | | "vars_list(E#Es) = vars E \<union> vars_list Es" | |
| 15172 | 282 | |
| 283 | ||
| 69597 | 284 | text\<open>Now prove the three equations for \<^term>\<open>vars\<close>\<close> | 
| 15172 | 285 | |
| 286 | lemma vars_Variable [simp]: "vars (Var N) = {N}"
 | |
| 287 | by (simp add: vars_def Var_def | |
| 288 | UN_equiv_class [OF equiv_exprel vars_respects]) | |
| 289 | ||
| 290 | lemma vars_Plus [simp]: "vars (Plus X Y) = vars X \<union> vars Y" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 291 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 292 |   have "\<And>U V. \<lbrakk>X = Abs_Exp (exprel``{U}); Y = Abs_Exp (exprel``{V})\<rbrakk>
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 293 | \<Longrightarrow> vars (Plus X Y) = vars X \<union> vars Y" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 294 | by (simp add: vars_def Plus UN_equiv_class [OF equiv_exprel vars_respects]) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 295 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 296 | by (meson eq_Abs_Exp) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 297 | qed | 
| 15172 | 298 | |
| 299 | lemma vars_FnCall [simp]: "vars (FnCall F Xs) = vars_list Xs" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 300 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 301 |   have "vars (Abs_Exp (exprel``{FNCALL F Us})) = vars_list (Abs_ExpList Us)" for Us
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 302 | by (induct Us) (auto simp: vars_def UN_equiv_class [OF equiv_exprel vars_respects]) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 303 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 304 | by (metis ExpList_rep FnCall) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 305 | qed | 
| 15172 | 306 | |
| 307 | lemma vars_FnCall_Nil: "vars (FnCall F Nil) = {}" 
 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 308 | by simp | 
| 15172 | 309 | |
| 310 | lemma vars_FnCall_Cons: "vars (FnCall F (X#Xs)) = vars X \<union> vars_list Xs" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 311 | by simp | 
| 15172 | 312 | |
| 313 | ||
| 60530 | 314 | subsection\<open>Injectivity Properties of Some Constructors\<close> | 
| 15172 | 315 | |
| 316 | lemma VAR_imp_eq: "VAR m \<sim> VAR n \<Longrightarrow> m = n" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 317 | by (drule exprel_imp_eq_freevars, simp) | 
| 15172 | 318 | |
| 69597 | 319 | text\<open>Can also be proved using the function \<^term>\<open>vars\<close>\<close> | 
| 15172 | 320 | lemma Var_Var_eq [iff]: "(Var m = Var n) = (m = n)" | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 321 | by (auto simp add: Var_def exprel_refl dest: VAR_imp_eq) | 
| 15172 | 322 | |
| 323 | lemma VAR_neqv_PLUS: "VAR m \<sim> PLUS X Y \<Longrightarrow> False" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 324 | using exprel_imp_eq_freediscrim by force | 
| 15172 | 325 | |
| 326 | theorem Var_neq_Plus [iff]: "Var N \<noteq> Plus X Y" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 327 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 328 |   have "\<And>U V. \<lbrakk>X = Abs_Exp (exprel``{U}); Y = Abs_Exp (exprel``{V})\<rbrakk> \<Longrightarrow> Var N \<noteq> Plus X Y"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 329 | using Plus VAR_neqv_PLUS Var_def by force | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 330 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 331 | by (meson eq_Abs_Exp) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 332 | qed | 
| 15172 | 333 | |
| 334 | theorem Var_neq_FnCall [iff]: "Var N \<noteq> FnCall F Xs" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 335 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 336 | have "\<And>Us. Var N \<noteq> FnCall F (Abs_ExpList Us)" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 337 | using FnCall Var_def exprel_imp_eq_freediscrim by fastforce | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 338 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 339 | by (metis ExpList_rep) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 340 | qed | 
| 15172 | 341 | |
| 69597 | 342 | subsection\<open>Injectivity of \<^term>\<open>FnCall\<close>\<close> | 
| 15172 | 343 | |
| 19736 | 344 | definition | 
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 345 | "fun" :: "exp \<Rightarrow> nat" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 346 |   where "fun X \<equiv> the_elem (\<Union>U \<in> Rep_Exp X. {freefun U})"
 | 
| 15172 | 347 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 348 | lemma fun_respects: "(\<lambda>U. {freefun U}) respects exprel"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 349 | by (auto simp add: congruent_def exprel_imp_eq_freefun) | 
| 15172 | 350 | |
| 351 | lemma fun_FnCall [simp]: "fun (FnCall F Xs) = F" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 352 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 353 | have "\<And>Us. fun (FnCall F (Abs_ExpList Us)) = F" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 354 | using FnCall UN_equiv_class [OF equiv_exprel] fun_def fun_respects by fastforce | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 355 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 356 | by (metis ExpList_rep) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 357 | qed | 
| 15172 | 358 | |
| 19736 | 359 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 360 | args :: "exp \<Rightarrow> exp list" where | 
| 39910 | 361 |   "args X = the_elem (\<Union>U \<in> Rep_Exp X. {Abs_ExpList (freeargs U)})"
 | 
| 15172 | 362 | |
| 60530 | 363 | text\<open>This result can probably be generalized to arbitrary equivalence | 
| 364 | relations, but with little benefit here.\<close> | |
| 15172 | 365 | lemma Abs_ExpList_eq: | 
| 366 | "(y, z) \<in> listrel exprel \<Longrightarrow> Abs_ExpList (y) = Abs_ExpList (z)" | |
| 18460 | 367 | by (induct set: listrel) simp_all | 
| 15172 | 368 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 369 | lemma args_respects: "(\<lambda>U. {Abs_ExpList (freeargs U)}) respects exprel"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 370 | by (auto simp add: congruent_def Abs_ExpList_eq exprel_imp_eqv_freeargs) | 
| 15172 | 371 | |
| 372 | lemma args_FnCall [simp]: "args (FnCall F Xs) = Xs" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 373 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 374 | have "\<And>Us. Xs = Abs_ExpList Us \<Longrightarrow> args (FnCall F Xs) = Xs" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 375 | by (simp add: FnCall args_def UN_equiv_class [OF equiv_exprel args_respects]) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 376 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 377 | by (metis ExpList_rep) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 378 | qed | 
| 15172 | 379 | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 380 | lemma FnCall_FnCall_eq [iff]: "(FnCall F Xs = FnCall F' Xs') \<longleftrightarrow> (F=F' \<and> Xs=Xs')" | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 381 | by (metis args_FnCall fun_FnCall) | 
| 15172 | 382 | |
| 383 | ||
| 60530 | 384 | subsection\<open>The Abstract Discriminator\<close> | 
| 63167 | 385 | text\<open>However, as \<open>FnCall_Var_neq_Var\<close> illustrates, we don't need this | 
| 60530 | 386 | function in order to prove discrimination theorems.\<close> | 
| 15172 | 387 | |
| 19736 | 388 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 389 | discrim :: "exp \<Rightarrow> int" where | 
| 39910 | 390 |   "discrim X = the_elem (\<Union>U \<in> Rep_Exp X. {freediscrim U})"
 | 
| 15172 | 391 | |
| 392 | lemma discrim_respects: "(\<lambda>U. {freediscrim U}) respects exprel"
 | |
| 40822 | 393 | by (auto simp add: congruent_def exprel_imp_eq_freediscrim) | 
| 15172 | 394 | |
| 69597 | 395 | text\<open>Now prove the four equations for \<^term>\<open>discrim\<close>\<close> | 
| 15172 | 396 | |
| 397 | lemma discrim_Var [simp]: "discrim (Var N) = 0" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 398 | by (simp add: discrim_def Var_def UN_equiv_class [OF equiv_exprel discrim_respects]) | 
| 15172 | 399 | |
| 400 | lemma discrim_Plus [simp]: "discrim (Plus X Y) = 1" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 401 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 402 |   have "\<And>U V. \<lbrakk>X = Abs_Exp (exprel``{U}); Y = Abs_Exp (exprel``{V})\<rbrakk> \<Longrightarrow> discrim (Plus X Y) = 1"
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 403 | by (simp add: discrim_def Plus UN_equiv_class [OF equiv_exprel discrim_respects]) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 404 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 405 | by (meson eq_Abs_Exp) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 406 | qed | 
| 15172 | 407 | |
| 408 | lemma discrim_FnCall [simp]: "discrim (FnCall F Xs) = 2" | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 409 | proof - | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 410 | have "discrim (FnCall F (Abs_ExpList Us)) = 2" for Us | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 411 | by (simp add: discrim_def FnCall UN_equiv_class [OF equiv_exprel discrim_respects]) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 412 | then show ?thesis | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 413 | by (metis ExpList_rep) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 414 | qed | 
| 15172 | 415 | |
| 60530 | 416 | text\<open>The structural induction rule for the abstract type\<close> | 
| 18460 | 417 | theorem exp_inducts: | 
| 15172 | 418 | assumes V: "\<And>nat. P1 (Var nat)" | 
| 419 | and P: "\<And>exp1 exp2. \<lbrakk>P1 exp1; P1 exp2\<rbrakk> \<Longrightarrow> P1 (Plus exp1 exp2)" | |
| 420 | and F: "\<And>nat list. P2 list \<Longrightarrow> P1 (FnCall nat list)" | |
| 421 | and Nil: "P2 []" | |
| 422 | and Cons: "\<And>exp list. \<lbrakk>P1 exp; P2 list\<rbrakk> \<Longrightarrow> P2 (exp # list)" | |
| 18460 | 423 | shows "P1 exp" and "P2 list" | 
| 424 | proof - | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 425 |   obtain U where exp: "exp = (Abs_Exp (exprel``{U}))" by (cases exp)
 | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 426 | obtain Us where list: "list = Abs_ExpList Us" by (metis ExpList_rep) | 
| 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 427 |   have "P1 (Abs_Exp (exprel``{U}))" and "P2 (Abs_ExpList Us)"
 | 
| 58305 
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
 blanchet parents: 
55417diff
changeset | 428 | proof (induct U and Us rule: compat_freeExp.induct compat_freeExp_list.induct) | 
| 18460 | 429 | case (VAR nat) | 
| 15172 | 430 | with V show ?case by (simp add: Var_def) | 
| 431 | next | |
| 432 | case (PLUS X Y) | |
| 75287 
7add2d5322a7
Tidied several ugly proofs in some elderly examples
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 433 |     with P [of "Abs_Exp (exprel``{X})" "Abs_Exp (exprel``{Y})"]
 | 
| 15172 | 434 | show ?case by (simp add: Plus) | 
| 435 | next | |
| 436 | case (FNCALL nat list) | |
| 437 | with F [of "Abs_ExpList list"] | |
| 438 | show ?case by (simp add: FnCall) | |
| 439 | next | |
| 440 | case Nil_freeExp | |
| 441 | with Nil show ?case by simp | |
| 442 | next | |
| 443 | case Cons_freeExp | |
| 18460 | 444 | with Cons show ?case by simp | 
| 15172 | 445 | qed | 
| 18460 | 446 | with exp and list show "P1 exp" and "P2 list" by (simp_all only:) | 
| 15172 | 447 | qed | 
| 448 | ||
| 449 | end |