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