src/HOL/ex/Simult.ML
author berghofe
Fri, 21 Jun 1996 12:18:50 +0200
changeset 1820 e381e1c51689
parent 1476 608483c2122a
child 1888 acb7363994cb
permissions -rw-r--r--
Classical tactics now use default claset.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     1
(*  Title:      HOL/ex/Simult.ML
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
Primitives for simultaneous recursive type definitions
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
  includes worked example of trees & forests
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     9
This is essentially the same data structure that on ex/term.ML, which is
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    10
simpler because it uses list as a new type former.  The approach in this
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    11
file may be superior for other simultaneous recursions.
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    12
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    13
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    14
open Simult;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    15
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    16
(*** Monotonicity and unfolding of the function ***)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    17
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    18
goal Simult.thy "mono(%Z.  A <*> Part Z In1 \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    19
\                      <+> ({Numb(0)} <+> Part Z In0 <*> Part Z In1))";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    20
by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    21
                      Part_mono] 1));
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    22
qed "TF_fun_mono";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    23
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    24
val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    25
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    26
goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    27
by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    28
qed "TF_mono";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    29
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    30
goalw Simult.thy [TF_def] "TF(sexp) <= sexp";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    31
by (rtac lfp_lowerbound 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    32
by (fast_tac (!claset addIs  sexp.intrs@[sexp_In0I, sexp_In1I]
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    33
                      addSEs [PartE]) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    34
qed "TF_sexp";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    35
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    36
(* A <= sexp ==> TF(A) <= sexp *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    37
val TF_subset_sexp = standard
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    38
    (TF_mono RS (TF_sexp RSN (2,subset_trans)));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    39
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    40
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    41
(** Elimination -- structural induction on the set TF **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    42
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    43
val TF_Rep_defs = [TCONS_def,FNIL_def,FCONS_def,NIL_def,CONS_def];
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    44
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    45
val major::prems = goalw Simult.thy TF_Rep_defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    46
 "[| i: TF(A);  \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    47
\    !!M N. [| M: A;  N: Part (TF A) In1;  R(N) |] ==> R(TCONS M N);    \
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    48
\    R(FNIL);                   \
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    49
\    !!M N. [| M:  Part (TF A) In0;  N: Part (TF A) In1;  R(M);  R(N) \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    50
\            |] ==> R(FCONS M N)    \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    51
\    |] ==> R(i)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    52
by (rtac ([TF_def, TF_fun_mono, major] MRS def_induct) 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    53
by (fast_tac (!claset addIs (prems@[PartI])
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    54
                       addEs [usumE, uprodE, PartE]) 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    55
qed "TF_induct";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    56
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    57
(*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    58
val prems = goalw Simult.thy [Part_def]
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    59
 "! M: TF(A). (M: Part (TF A) In0 --> P(M)) & (M: Part (TF A) In1 --> Q(M)) \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    60
\ ==> (! M: Part (TF A) In0. P(M)) & (! M: Part (TF A) In1. Q(M))";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    61
by (cfast_tac prems 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    62
qed "TF_induct_lemma";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    63
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    64
AddSIs [PartI];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    65
AddSDs [In0_inject, In1_inject];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    66
AddSEs [In0_neq_In1, In1_neq_In0, PartE];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    67
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    68
(*Could prove  ~ TCONS M N : Part (TF A) In1  etc. *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    69
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    70
(*Induction on TF with separate predicates P, Q*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    71
val prems = goalw Simult.thy TF_Rep_defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    72
    "[| !!M N. [| M: A;  N: Part (TF A) In1;  Q(N) |] ==> P(TCONS M N); \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    73
\       Q(FNIL);        \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    74
\       !!M N. [| M:  Part (TF A) In0;  N: Part (TF A) In1;  P(M);  Q(N) \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    75
\               |] ==> Q(FCONS M N)     \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    76
\    |] ==> (! M: Part (TF A) In0. P(M)) & (! N: Part (TF A) In1. Q(N))";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    77
by (rtac (ballI RS TF_induct_lemma) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    78
by (etac TF_induct 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    79
by (rewrite_goals_tac TF_Rep_defs);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    80
by (ALLGOALS (fast_tac (!claset addIs prems)));
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    81
(*29 secs??*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    82
qed "Tree_Forest_induct";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    83
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    84
(*Induction for the abstract types 'a tree, 'a forest*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    85
val prems = goalw Simult.thy [Tcons_def,Fnil_def,Fcons_def]
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    86
    "[| !!x ts. Q(ts) ==> P(Tcons x ts);     \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    87
\       Q(Fnil);        \
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    88
\       !!t ts. [| P(t);  Q(ts) |] ==> Q(Fcons t ts)    \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    89
\    |] ==> (! t. P(t)) & (! ts. Q(ts))";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    90
by (res_inst_tac [("P1","%z.P(Abs_Tree(z))"),
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    91
                  ("Q1","%z.Q(Abs_Forest(z))")] 
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    92
    (Tree_Forest_induct RS conjE) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    93
(*Instantiates ?A1 to range(Leaf). *)
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    94
by (fast_tac (!claset addSEs [Rep_Tree_inverse RS subst, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    95
                             Rep_Forest_inverse RS subst] 
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    96
                     addSIs [Rep_Tree,Rep_Forest]) 4);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    97
(*Cannot use simplifier: the rewrites work in the wrong direction!*)
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
    98
by (ALLGOALS (fast_tac (!claset addSEs [Abs_Tree_inverse RS subst,
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    99
                          Abs_Forest_inverse RS subst] 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   100
                     addSIs prems)));
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   101
qed "tree_forest_induct";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   102
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   103
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   104
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   105
(*** Isomorphisms ***)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   106
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   107
goal Simult.thy "inj(Rep_Tree)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   108
by (rtac inj_inverseI 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   109
by (rtac Rep_Tree_inverse 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   110
qed "inj_Rep_Tree";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   111
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   112
goal Simult.thy "inj_onto Abs_Tree (Part (TF(range Leaf)) In0)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   113
by (rtac inj_onto_inverseI 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   114
by (etac Abs_Tree_inverse 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   115
qed "inj_onto_Abs_Tree";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   116
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   117
goal Simult.thy "inj(Rep_Forest)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   118
by (rtac inj_inverseI 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   119
by (rtac Rep_Forest_inverse 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   120
qed "inj_Rep_Forest";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   121
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   122
goal Simult.thy "inj_onto Abs_Forest (Part (TF(range Leaf)) In1)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   123
by (rtac inj_onto_inverseI 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   124
by (etac Abs_Forest_inverse 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   125
qed "inj_onto_Abs_Forest";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   126
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   127
(** Introduction rules for constructors **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   128
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   129
(* c : A <*> Part (TF A) In1 
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   130
        <+> {Numb(0)} <+> Part (TF A) In0 <*> Part (TF A) In1 ==> c : TF(A) *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   131
val TF_I = TF_unfold RS equalityD2 RS subsetD;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   132
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   133
(*For reasoning about the representation*)
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   134
AddIs [TF_I, uprodI, usum_In0I, usum_In1I];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   135
AddSEs [Scons_inject];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   136
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   137
val prems = goalw Simult.thy TF_Rep_defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   138
    "[| a: A;  M: Part (TF A) In1 |] ==> TCONS a M : Part (TF A) In0";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   139
by (fast_tac (!claset addIs prems) 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   140
qed "TCONS_I";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   141
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   142
(* FNIL is a TF(A) -- this also justifies the type definition*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   143
goalw Simult.thy TF_Rep_defs "FNIL: Part (TF A) In1";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   144
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   145
qed "FNIL_I";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   146
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   147
val prems = goalw Simult.thy TF_Rep_defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   148
    "[| M: Part (TF A) In0;  N: Part (TF A) In1 |] ==> \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   149
\    FCONS M N : Part (TF A) In1";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   150
by (fast_tac (!claset addIs prems) 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   151
qed "FCONS_I";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   152
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   153
(** Injectiveness of TCONS and FCONS **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   154
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   155
goalw Simult.thy TF_Rep_defs "(TCONS K M=TCONS L N) = (K=L & M=N)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   156
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   157
qed "TCONS_TCONS_eq";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   158
bind_thm ("TCONS_inject", (TCONS_TCONS_eq RS iffD1 RS conjE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   159
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   160
goalw Simult.thy TF_Rep_defs "(FCONS K M=FCONS L N) = (K=L & M=N)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   161
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   162
qed "FCONS_FCONS_eq";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   163
bind_thm ("FCONS_inject", (FCONS_FCONS_eq RS iffD1 RS conjE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   164
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   165
(** Distinctness of TCONS, FNIL and FCONS **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   166
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   167
goalw Simult.thy TF_Rep_defs "TCONS M N ~= FNIL";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   168
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   169
qed "TCONS_not_FNIL";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   170
bind_thm ("FNIL_not_TCONS", (TCONS_not_FNIL RS not_sym));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   171
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   172
bind_thm ("TCONS_neq_FNIL", (TCONS_not_FNIL RS notE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   173
val FNIL_neq_TCONS = sym RS TCONS_neq_FNIL;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   174
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   175
goalw Simult.thy TF_Rep_defs "FCONS M N ~= FNIL";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   176
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   177
qed "FCONS_not_FNIL";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   178
bind_thm ("FNIL_not_FCONS", (FCONS_not_FNIL RS not_sym));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   179
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   180
bind_thm ("FCONS_neq_FNIL", (FCONS_not_FNIL RS notE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   181
val FNIL_neq_FCONS = sym RS FCONS_neq_FNIL;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   182
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   183
goalw Simult.thy TF_Rep_defs "TCONS M N ~= FCONS K L";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   184
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   185
qed "TCONS_not_FCONS";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   186
bind_thm ("FCONS_not_TCONS", (TCONS_not_FCONS RS not_sym));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   187
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   188
bind_thm ("TCONS_neq_FCONS", (TCONS_not_FCONS RS notE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   189
val FCONS_neq_TCONS = sym RS TCONS_neq_FCONS;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   190
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   191
(*???? Too many derived rules ????
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   192
  Automatically generate symmetric forms?  Always expand TF_Rep_defs? *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   193
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   194
(** Injectiveness of Tcons and Fcons **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   195
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   196
(*For reasoning about abstract constructors*)
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   197
AddSIs [Rep_Tree, Rep_Forest, TCONS_I, FNIL_I, FCONS_I];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   198
AddSEs [TCONS_inject, FCONS_inject,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   199
                           TCONS_neq_FNIL, FNIL_neq_TCONS,
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   200
                           FCONS_neq_FNIL, FNIL_neq_FCONS,
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   201
                           TCONS_neq_FCONS, FCONS_neq_TCONS];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   202
AddSDs [inj_onto_Abs_Tree RS inj_ontoD,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   203
                           inj_onto_Abs_Forest RS inj_ontoD,
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   204
                           inj_Rep_Tree RS injD, inj_Rep_Forest RS injD,
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   205
                           Leaf_inject];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   206
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   207
goalw Simult.thy [Tcons_def] "(Tcons x xs=Tcons y ys) = (x=y & xs=ys)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   208
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   209
qed "Tcons_Tcons_eq";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   210
bind_thm ("Tcons_inject", (Tcons_Tcons_eq RS iffD1 RS conjE));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   211
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   212
goalw Simult.thy [Fcons_def,Fnil_def] "Fcons x xs ~= Fnil";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   213
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   214
qed "Fcons_not_Fnil";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   215
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   216
bind_thm ("Fcons_neq_Fnil", Fcons_not_Fnil RS notE);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   217
val Fnil_neq_Fcons = sym RS Fcons_neq_Fnil;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   218
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   219
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   220
(** Injectiveness of Fcons **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   221
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   222
goalw Simult.thy [Fcons_def] "(Fcons x xs=Fcons y ys) = (x=y & xs=ys)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1476
diff changeset
   223
by (Fast_tac 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   224
qed "Fcons_Fcons_eq";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   225
bind_thm ("Fcons_inject", Fcons_Fcons_eq RS iffD1 RS conjE);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   226
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   227
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   228
(*** TF_rec -- by wf recursion on pred_sexp ***)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   229
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   230
goal Simult.thy
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   231
   "(%M. TF_rec M b c d) = wfrec (trancl pred_sexp) \
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   232
                         \       (%g. Case (Split(%x y. b x y (g y))) \
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   233
                         \           (List_case c (%x y. d x y (g x) (g y))))";
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   234
by (simp_tac (HOL_ss addsimps [TF_rec_def]) 1);
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   235
val TF_rec_unfold = (wf_pred_sexp RS wf_trancl) RS 
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   236
                    ((result() RS eq_reflection) RS def_wfrec);
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   237
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   238
(*---------------------------------------------------------------------------
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   239
 * Old: 
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   240
 * val TF_rec_unfold =
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   241
 *   wf_pred_sexp RS wf_trancl RS (TF_rec_def RS def_wfrec);
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1465
diff changeset
   242
 *---------------------------------------------------------------------------*)
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   243
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   244
(** conversion rules for TF_rec **)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   245
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   246
goalw Simult.thy [TCONS_def]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   247
    "!!M N. [| M: sexp;  N: sexp |] ==>         \
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   248
\           TF_rec (TCONS M N) b c d = b M N (TF_rec N b c d)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   249
by (rtac (TF_rec_unfold RS trans) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   250
by (simp_tac (!simpset addsimps [Case_In0, Split]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   251
by (asm_simp_tac (!simpset addsimps [In0_def]) 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   252
qed "TF_rec_TCONS";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   253
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   254
goalw Simult.thy [FNIL_def] "TF_rec FNIL b c d = c";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   255
by (rtac (TF_rec_unfold RS trans) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   256
by (simp_tac (HOL_ss addsimps [Case_In1, List_case_NIL]) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   257
qed "TF_rec_FNIL";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   258
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   259
goalw Simult.thy [FCONS_def]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   260
 "!!M N. [| M: sexp;  N: sexp |] ==>    \
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   261
\        TF_rec (FCONS M N) b c d = d M N (TF_rec M b c d) (TF_rec N b c d)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   262
by (rtac (TF_rec_unfold RS trans) 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   263
by (simp_tac (HOL_ss addsimps [Case_In1, List_case_CONS]) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   264
by (asm_simp_tac (!simpset addsimps [CONS_def,In1_def]) 1);
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   265
qed "TF_rec_FCONS";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   266
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   267
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   268
(*** tree_rec, forest_rec -- by TF_rec ***)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   269
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   270
val Rep_Tree_in_sexp =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   271
    [range_Leaf_subset_sexp RS TF_subset_sexp RS (Part_subset RS subset_trans),
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   272
     Rep_Tree] MRS subsetD;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   273
val Rep_Forest_in_sexp =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   274
    [range_Leaf_subset_sexp RS TF_subset_sexp RS (Part_subset RS subset_trans),
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   275
     Rep_Forest] MRS subsetD;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   276
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   277
val tf_rec_ss = HOL_ss addsimps
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   278
  [TF_rec_TCONS, TF_rec_FNIL, TF_rec_FCONS,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   279
   TCONS_I, FNIL_I, FCONS_I, Rep_Tree, Rep_Forest,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   280
   Rep_Tree_inverse, Rep_Forest_inverse,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   281
   Abs_Tree_inverse, Abs_Forest_inverse,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   282
   inj_Leaf, Inv_f_f, sexp.LeafI, range_eqI,
3ae9fe3c0f68 added local simpsets
clasohm
parents: 969
diff changeset
   283
   Rep_Tree_in_sexp, Rep_Forest_in_sexp];
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   284
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   285
goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def]
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   286
    "tree_rec (Tcons a tf) b c d = b a tf (forest_rec tf b c d)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   287
by (simp_tac tf_rec_ss 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   288
qed "tree_rec_Tcons";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   289
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   290
goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec Fnil b c d = c";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   291
by (simp_tac tf_rec_ss 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   292
qed "forest_rec_Fnil";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   293
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   294
goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def]
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   295
    "forest_rec (Fcons t tf) b c d = \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   296
\    d t tf (tree_rec t b c d) (forest_rec tf b c d)";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   297
by (simp_tac tf_rec_ss 1);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   298
qed "forest_rec_Cons";