src/HOL/Induct/Sexp.ML
author wenzelm
Mon, 08 May 2000 20:59:30 +0200
changeset 8840 18b76c137c41
child 11481 c77e5401f2ff
permissions -rw-r--r--
moved theory Sexp to Induct examples;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8840
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Sexp
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     5
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     6
S-expressions, general binary trees for defining recursive data structures
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     7
*)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     8
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
     9
open Sexp;
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    10
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    11
(** sexp_case **)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    12
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    13
Goalw [sexp_case_def] "sexp_case c d e (Leaf a) = c(a)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    14
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    15
qed "sexp_case_Leaf";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    16
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    17
Goalw [sexp_case_def] "sexp_case c d e (Numb k) = d(k)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    18
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    19
qed "sexp_case_Numb";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    20
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    21
Goalw [sexp_case_def] "sexp_case c d e (Scons M N) = e M N";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    22
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    23
qed "sexp_case_Scons";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    24
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    25
Addsimps [sexp_case_Leaf, sexp_case_Numb, sexp_case_Scons];
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    26
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    27
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    28
(** Introduction rules for sexp constructors **)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    29
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    30
val [prem] = goalw Sexp.thy [In0_def] "M: sexp ==> In0(M) : sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    31
by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    32
qed "sexp_In0I";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    33
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    34
val [prem] = goalw Sexp.thy [In1_def] "M: sexp ==> In1(M) : sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    35
by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    36
qed "sexp_In1I";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    37
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    38
AddIs sexp.intrs;
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    39
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    40
Goal "range(Leaf) <= sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    41
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    42
qed "range_Leaf_subset_sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    43
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    44
val [major] = goal Sexp.thy "Scons M N : sexp ==> M: sexp & N: sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    45
by (rtac (major RS setup_induction) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    46
by (etac sexp.induct 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    47
by (ALLGOALS Blast_tac);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    48
qed "Scons_D";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    49
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    50
(** Introduction rules for 'pred_sexp' **)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    51
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    52
Goalw [pred_sexp_def] "pred_sexp <= sexp <*> sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    53
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    54
qed "pred_sexp_subset_Sigma";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    55
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    56
(* (a,b) : pred_sexp^+ ==> a : sexp *)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    57
val trancl_pred_sexpD1 = 
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    58
    pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD1
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    59
and trancl_pred_sexpD2 = 
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    60
    pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD2;
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    61
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    62
Goalw [pred_sexp_def]
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    63
    "!!M. [| M: sexp;  N: sexp |] ==> (M, Scons M N) : pred_sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    64
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    65
qed "pred_sexpI1";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    66
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    67
Goalw [pred_sexp_def]
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    68
    "!!M. [| M: sexp;  N: sexp |] ==> (N, Scons M N) : pred_sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    69
by (Blast_tac 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    70
qed "pred_sexpI2";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    71
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    72
(*Combinations involving transitivity and the rules above*)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    73
val pred_sexp_t1 = pred_sexpI1 RS r_into_trancl
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    74
and pred_sexp_t2 = pred_sexpI2 RS r_into_trancl;
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    75
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    76
val pred_sexp_trans1 = pred_sexp_t1 RSN (2, trans_trancl RS transD)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    77
and pred_sexp_trans2 = pred_sexp_t2 RSN (2, trans_trancl RS transD);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    78
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    79
(*Proves goals of the form (M,N):pred_sexp^+ provided M,N:sexp*)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    80
Addsimps (sexp.intrs @ [pred_sexp_t1, pred_sexp_t2,
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    81
                        pred_sexp_trans1, pred_sexp_trans2, cut_apply]);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    82
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    83
val major::prems = goalw Sexp.thy [pred_sexp_def]
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    84
    "[| p : pred_sexp;                                       \
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    85
\       !!M N. [| p = (M, Scons M N);  M: sexp;  N: sexp |] ==> R; \
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    86
\       !!M N. [| p = (N, Scons M N);  M: sexp;  N: sexp |] ==> R  \
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    87
\    |] ==> R";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    88
by (cut_facts_tac [major] 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    89
by (REPEAT (eresolve_tac ([asm_rl,emptyE,insertE,UN_E]@prems) 1));
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    90
qed "pred_sexpE";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    91
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    92
Goal "wf(pred_sexp)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    93
by (rtac (pred_sexp_subset_Sigma RS wfI) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    94
by (etac sexp.induct 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    95
by (ALLGOALS (blast_tac (claset() addSEs [allE, pred_sexpE])));
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    96
qed "wf_pred_sexp";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    97
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    98
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
    99
(*** sexp_rec -- by wf recursion on pred_sexp ***)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   100
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   101
Goal "(%M. sexp_rec M c d e) = wfrec pred_sexp \
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   102
                       \ (%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2)))";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   103
by (simp_tac (HOL_ss addsimps [sexp_rec_def]) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   104
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   105
(* sexp_rec a c d e =
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   106
   sexp_case c d
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   107
    (%N1 N2.
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   108
        e N1 N2 (cut (%M. sexp_rec M c d e) pred_sexp a N1)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   109
         (cut (%M. sexp_rec M c d e) pred_sexp a N2)) a
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   110
*)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   111
bind_thm("sexp_rec_unfold", 
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   112
	 [result() RS eq_reflection, wf_pred_sexp] MRS def_wfrec);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   113
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   114
(** conversion rules **)
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   115
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   116
Goal "sexp_rec (Leaf a) c d h = c(a)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   117
by (stac sexp_rec_unfold 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   118
by (rtac sexp_case_Leaf 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   119
qed "sexp_rec_Leaf";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   120
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   121
Goal "sexp_rec (Numb k) c d h = d(k)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   122
by (stac sexp_rec_unfold 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   123
by (rtac sexp_case_Numb 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   124
qed "sexp_rec_Numb";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   125
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   126
Goal "[| M: sexp;  N: sexp |] ==> \
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   127
\    sexp_rec (Scons M N) c d h = h M N (sexp_rec M c d h) (sexp_rec N c d h)";
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   128
by (rtac (sexp_rec_unfold RS trans) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   129
by (asm_simp_tac (simpset() addsimps [pred_sexpI1, pred_sexpI2]) 1);
18b76c137c41 moved theory Sexp to Induct examples;
wenzelm
parents:
diff changeset
   130
qed "sexp_rec_Scons";