src/HOL/Induct/Term.ML
author paulson
Thu, 11 Nov 1999 10:25:17 +0100
changeset 8005 b64d86018785
parent 7847 5a3fa0c4b215
permissions -rw-r--r--
new-style infix declaration for "image"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
     1
(*  Title:      HOL/Induct/Term.ML
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     2
    ID:         $Id$
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
     4
    Copyright   1998  TU Muenchen
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     5
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
     6
Terms over a given alphabet -- function applications
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
     9
(** a simple theorem about composition of substitutions **)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    10
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    11
Goal "(subst_term ((subst_term f1) o f2) t) =  \
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    12
  \  (subst_term f1 (subst_term f2 t)) &  \
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    13
  \  (subst_term_list ((subst_term f1) o f2) ts) =  \
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    14
  \  (subst_term_list f1 (subst_term_list f2 ts))";
7847
5a3fa0c4b215 Eliminated mutual_induct_tac.
berghofe
parents: 5738
diff changeset
    15
by (induct_tac "t ts" 1);
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    16
by (ALLGOALS Asm_full_simp_tac);
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    17
qed "subst_comp";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    18
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    19
(**** alternative induction rule ****)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    20
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    21
bind_thm ("term_induct2", prove_goal thy
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    22
  "[| !!v. P (Var v);  \
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    23
   \  !!f ts. list_all P ts ==> P (App f ts) |] ==>  \
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    24
   \    P t & list_all P ts"
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    25
  (fn prems =>
7847
5a3fa0c4b215 Eliminated mutual_induct_tac.
berghofe
parents: 5738
diff changeset
    26
    [induct_tac "t ts" 1,
5738
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    27
     resolve_tac prems 1,
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    28
     resolve_tac prems 1,
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    29
     atac 1,
0d8698c15439 Terms are now defined using the new datatype package.
berghofe
parents: 5278
diff changeset
    30
     ALLGOALS Asm_simp_tac]) RS conjunct1);