src/ZF/ex/Term.ML
author paulson
Wed, 13 Jan 1999 11:57:09 +0100
changeset 6112 5e4871c5136b
parent 6046 2c8a8be36c94
child 6153 bff90585cce5
permissions -rw-r--r--
datatype package improvements
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/ex/Term.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
486
6b58082796f6 Misc minor updates
lcp
parents: 477
diff changeset
     4
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Datatype definition of terms over an alphabet.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
Illustrates the list functor (essentially the same type as in Trees & Forests)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    10
Goal "term(A) = A * list(term(A))";
529
f0d16216e394 ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
lcp
parents: 515
diff changeset
    11
let open term;  val rew = rewrite_rule con_defs in  
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    12
by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
529
f0d16216e394 ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
lcp
parents: 515
diff changeset
    13
end;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    14
qed "term_unfold";
434
89d45187f04d Various updates and tidying
lcp
parents: 71
diff changeset
    15
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
(*Induction on term(A) followed by induction on List *)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    17
val major::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
    "[| t: term(A);  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
\       !!x.      [| x: A |] ==> P(Apply(x,Nil));  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
\       !!x z zs. [| x: A;  z: term(A);  zs: list(term(A));  P(Apply(x,zs))  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
\                 |] ==> P(Apply(x, Cons(z,zs)))  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
\    |] ==> P(t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    23
by (rtac (major RS term.induct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    24
by (etac list.induct 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
by (etac CollectE 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
by (REPEAT (ares_tac (prems@[list_CollectD]) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    27
qed "term_induct2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
(*Induction on term(A) to prove an equation*)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    30
val major::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
    "[| t: term(A);  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
\       !!x zs. [| x: A;  zs: list(term(A));  map(f,zs) = map(g,zs) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
\               f(Apply(x,zs)) = g(Apply(x,zs))  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
\    |] ==> f(t)=g(t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    35
by (rtac (major RS term.induct) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
by (resolve_tac prems 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
    38
qed "term_induct_eqn";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
(**  Lemmas to justify using "term" in other recursive type definitions **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    42
Goalw term.defs "A<=B ==> term(A) <= term(B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
by (rtac lfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    44
by (REPEAT (rtac term.bnd_mono 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    46
qed "term_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
(*Easily provable by induction also*)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    49
Goalw (term.defs@term.con_defs) "term(univ(A)) <= univ(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
by (rtac lfp_lowerbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
by (rtac (A_subset_univ RS univ_mono) 2);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    52
by Safe_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    54
qed "term_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
434
89d45187f04d Various updates and tidying
lcp
parents: 71
diff changeset
    56
val term_subset_univ = 
89d45187f04d Various updates and tidying
lcp
parents: 71
diff changeset
    57
    term_mono RS (term_univ RSN (2,subset_trans)) |> standard;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    59
Goal "[| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
434
89d45187f04d Various updates and tidying
lcp
parents: 71
diff changeset
    60
by (REPEAT (ares_tac [term_subset_univ RS subsetD] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    61
qed "term_into_univ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    62
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    63
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    64
(*** term_rec -- by Vset recursion ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    65
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    66
(*Lemma: map works correctly on the underlying list of terms*)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    67
Goal "[| l: list(A);  Ord(i) |] ==>  \
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    68
\     rank(l)<i --> map(%z. (lam x:Vset(i).h(x)) ` z, l) = map(h,l)";
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    69
by (etac list.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2034
diff changeset
    70
by (Simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    71
by (rtac impI 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    72
by (subgoal_tac "rank(a)<i & rank(l) < i" 1);
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    73
by (asm_simp_tac (simpset() addsimps [VsetI]) 1);
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    74
by (full_simp_tac (simpset() addsimps list.con_defs) 1);
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    75
by (blast_tac (claset() addDs (rank_rls RL [lt_trans])) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    76
qed "map_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    77
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    78
(*Typing premise is necessary to invoke map_lemma*)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    79
Goal "ts: list(A) ==> \
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    80
\     term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    81
by (rtac (term_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    82
by (rewrite_goals_tac term.con_defs);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    83
by (asm_simp_tac (simpset() addsimps [Ord_rank, rank_pair2, map_lemma]) 1);;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
    84
qed "term_rec";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    85
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    86
(*Slightly odd typing condition on r in the second premise!*)
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
    87
val major::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    88
    "[| t: term(A);                                     \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    89
\       !!x zs r. [| x: A;  zs: list(term(A));          \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    90
\                    r: list(UN t:term(A). C(t)) |]     \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    91
\                 ==> d(x, zs, r): C(Apply(x,zs))       \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    92
\    |] ==> term_rec(t,d) : C(t)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    93
by (rtac (major RS term.induct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    94
by (forward_tac [list_CollectD] 1);
2034
5079fdf938dd Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    95
by (stac term_rec 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    96
by (REPEAT (ares_tac prems 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    97
by (etac list.induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    98
by (ALLGOALS (asm_simp_tac (simpset() addsimps [term_rec])));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    99
by (etac CollectE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   100
by (REPEAT (ares_tac [list.Cons_I, UN_I] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   101
qed "term_rec_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   102
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
   103
val [rew,tslist] = Goal
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   104
    "[| !!t. j(t)==term_rec(t,d);  ts: list(A) |] ==> \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 2637
diff changeset
   105
\    j(Apply(a,ts)) = d(a, ts, map(%Z. j(Z), ts))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   106
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   107
by (rtac (tslist RS term_rec) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
   108
qed "def_term_rec";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   109
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
   110
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   111
    "[| t: term(A);                                          \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   112
\       !!x zs r. [| x: A;  zs: list(term(A));  r: list(C) |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   113
\                 ==> d(x, zs, r): C                 \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   114
\    |] ==> term_rec(t,d) : C";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   115
by (REPEAT (ares_tac (term_rec_type::prems) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   116
by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
   117
qed "term_rec_simple_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   118
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   119
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   120
(** term_map **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   121
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   122
bind_thm ("term_map", term_map_def RS def_term_rec);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   123
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
   124
val prems = Goalw [term_map_def]
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   125
    "[| t: term(A);  !!x. x: A ==> f(x): B |] ==> term_map(f,t) : term(B)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   126
by (REPEAT (ares_tac ([term_rec_simple_type, term.Apply_I] @ prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
   127
qed "term_map_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   128
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
   129
Goal "t: term(A) ==> term_map(f,t) : term({f(u). u:A})";
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5147
diff changeset
   130
by (etac term_map_type 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   131
by (etac RepFunI 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 529
diff changeset
   132
qed "term_map_type2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   133
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   134
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   135
(** term_size **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   136
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   137
bind_thm ("term_size", term_size_def RS def_term_rec);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   138
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   139
Goalw [term_size_def] "t: term(A) ==> term_size(t) : nat";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   140
by (REPEAT (ares_tac [term_rec_simple_type, list_add_type, nat_succI] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   141
qed "term_size_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   142
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   143
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   144
(** reflect **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   145
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   146
bind_thm ("reflect", reflect_def RS def_term_rec);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   147
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   148
Goalw [reflect_def] "t: term(A) ==> reflect(t) : term(A)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   149
by (REPEAT (ares_tac [term_rec_simple_type, rev_type, term.Apply_I] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   150
qed "reflect_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   151
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   152
(** preorder **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   153
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   154
bind_thm ("preorder", preorder_def RS def_term_rec);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   155
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   156
Goalw [preorder_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   157
    "t: term(A) ==> preorder(t) : list(A)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   158
by (REPEAT (ares_tac [term_rec_simple_type, list.Cons_I, flat_type] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   159
qed "preorder_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   160
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   161
(** postorder **)
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   162
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   163
bind_thm ("postorder", postorder_def RS def_term_rec);
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   164
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   165
Goalw [postorder_def]
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   166
    "t: term(A) ==> postorder(t) : list(A)";
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   167
by (REPEAT (ares_tac [term_rec_simple_type] 1));
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   168
by (Asm_simp_tac 1);
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   169
qed "postorder_type";
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   170
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   171
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   172
(** Term simplification **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   173
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   174
val term_typechecks =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   175
    [term.Apply_I, term_map_type, term_map_type2, term_size_type, 
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   176
     reflect_type, preorder_type, postorder_type];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   177
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   178
(*map_type2 and term_map_type2 instantiate variables*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   179
simpset_ref() := simpset()
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   180
      addsimps [term_rec, term_map, term_size, reflect, preorder, postorder]
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2496
diff changeset
   181
      setSolver type_auto_tac (list_typechecks@term_typechecks);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   182
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   183
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   184
(** theorems about term_map **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   185
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   186
Addsimps [thm "List.map_compose"];  (*there is also TF.map_compose*)
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   187
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   188
Goal "t: term(A) ==> term_map(%u. u, t) = t";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   189
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   190
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   191
qed "term_map_ident";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   192
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   193
Goal "t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u. f(g(u)), t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   194
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   195
by (asm_simp_tac (simpset() addsimps []) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   196
qed "term_map_compose";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   197
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   198
Goal "t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   199
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   200
by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   201
qed "term_map_reflect";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   202
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   203
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   204
(** theorems about term_size **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   205
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   206
Goal "t: term(A) ==> term_size(term_map(f,t)) = term_size(t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   207
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   208
by (asm_simp_tac (simpset() addsimps []) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   209
qed "term_size_term_map";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   210
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   211
Goal "t: term(A) ==> term_size(reflect(t)) = term_size(t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   212
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   213
by (asm_simp_tac(simpset() addsimps [rev_map_distrib RS sym, list_add_rev]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   214
qed "term_size_reflect";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   215
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   216
Goal "t: term(A) ==> term_size(t) = length(preorder(t))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   217
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   218
by (asm_simp_tac (simpset() addsimps [length_flat]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   219
qed "term_size_length";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   220
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   221
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   222
(** theorems about reflect **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   223
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   224
Goal "t: term(A) ==> reflect(reflect(t)) = t";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   225
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   226
by (asm_simp_tac (simpset() addsimps [rev_map_distrib]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   227
qed "reflect_reflect_ident";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   228
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   229
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   230
(** theorems about preorder **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   231
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   232
Goal "t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   233
by (etac term_induct_eqn 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   234
by (asm_simp_tac (simpset() addsimps [map_flat]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   235
qed "preorder_term_map";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   236
6112
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   237
Goal "t: term(A) ==> preorder(reflect(t)) = rev(postorder(t))";
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   238
by (etac term_induct_eqn 1);
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   239
by (asm_simp_tac(simpset() addsimps [rev_app_distrib, rev_flat, 
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   240
				     rev_map_distrib RS sym]) 1);
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   241
qed "preorder_reflect_eq_rev_postorder";
5e4871c5136b datatype package improvements
paulson
parents: 6046
diff changeset
   242
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   243
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   244
writeln"Reached end of file.";