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