src/HOL/Isar_examples/NestedDatatype.thy
author wenzelm
Mon, 16 Mar 2009 18:24:30 +0100
changeset 30549 d2d7874648bd
parent 23373 ead82c82da9e
permissions -rw-r--r--
simplified method setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
     1
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
     2
(* $Id$ *)
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
     3
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
     4
header {* Nested datatypes *}
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
     5
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 11809
diff changeset
     6
theory NestedDatatype imports Main begin
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
     7
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
     8
subsection {* Terms and substitution *}
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
     9
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    10
datatype ('a, 'b) "term" =
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    11
    Var 'a
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    12
  | App 'b "('a, 'b) term list"
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    13
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    14
consts
8717
20c42415c07d plain ASCII;
wenzelm
parents: 8676
diff changeset
    15
  subst_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term"
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    16
  subst_term_list ::
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    17
    "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list"
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    18
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    19
primrec (subst)
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    20
  "subst_term f (Var a) = f a"
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    21
  "subst_term f (App b ts) = App b (subst_term_list f ts)"
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    22
  "subst_term_list f [] = []"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    23
  "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    24
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    25
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    26
text {*
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    27
 \medskip A simple lemma about composition of substitutions.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    28
*}
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    29
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18153
diff changeset
    30
lemma "subst_term (subst_term f1 o f2) t =
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18153
diff changeset
    31
      subst_term f1 (subst_term f2 t)"
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18153
diff changeset
    32
  and "subst_term_list (subst_term f1 o f2) ts =
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    33
      subst_term_list f1 (subst_term_list f2 ts)"
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10458
diff changeset
    34
  by (induct t and ts) simp_all
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    35
9659
wenzelm
parents: 8717
diff changeset
    36
lemma "subst_term (subst_term f1 o f2) t =
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    37
  subst_term f1 (subst_term f2 t)"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    38
proof -
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    39
  let "?P t" = ?thesis
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    40
  let ?Q = "\<lambda>ts. subst_term_list (subst_term f1 o f2) ts =
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    41
    subst_term_list f1 (subst_term_list f2 ts)"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    42
  show ?thesis
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    43
  proof (induct t)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    44
    fix a show "?P (Var a)" by simp
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    45
  next
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    46
    fix b ts assume "?Q ts"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    47
    then show "?P (App b ts)"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    48
      by (simp only: subst.simps)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    49
  next
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    50
    show "?Q []" by simp
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    51
  next
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    52
    fix t ts
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    53
    assume "?P t" "?Q ts" then show "?Q (t # ts)"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    54
      by (simp only: subst.simps)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    55
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    56
qed
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    57
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    58
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    59
subsection {* Alternative induction *}
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    60
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    61
theorem term_induct' [case_names Var App]:
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    62
  assumes var: "!!a. P (Var a)"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    63
    and app: "!!b ts. list_all P ts ==> P (App b ts)"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    64
  shows "P t"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    65
proof (induct t)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    66
  fix a show "P (Var a)" by (rule var)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    67
next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    68
  fix b t ts assume "list_all P ts"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    69
  then show "P (App b ts)" by (rule app)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    70
next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    71
  show "list_all P []" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    72
next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    73
  fix t ts assume "P t" "list_all P ts"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    74
  then show "list_all P (t # ts)" by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    75
qed
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    76
8717
20c42415c07d plain ASCII;
wenzelm
parents: 8676
diff changeset
    77
lemma
20c42415c07d plain ASCII;
wenzelm
parents: 8676
diff changeset
    78
  "subst_term (subst_term f1 o f2) t = subst_term f1 (subst_term f2 t)"
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10458
diff changeset
    79
proof (induct t rule: term_induct')
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10458
diff changeset
    80
  case (Var a)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
    81
  show ?case by (simp add: o_def)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    82
next
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10458
diff changeset
    83
  case (App b ts)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18460
diff changeset
    84
  then show ?case by (induct ts) simp_all
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    85
qed
8676
4bf18b611a75 added NestedDatatype.thy;
wenzelm
parents:
diff changeset
    86
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    87
end