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