src/ZF/ex/Ntree.ML
author wenzelm
Sat, 20 Oct 2001 20:20:41 +0200
changeset 11852 a528a716a312
parent 11316 b4e71bd751e4
permissions -rw-r--r--
added TextIO.stdErr;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/ex/Ntree.ML
486
6b58082796f6 Misc minor updates
lcp
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:
diff changeset
     4
    Copyright   1994  University of Cambridge
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
     5
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
     6
Datatype definition n-ary branching trees
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
     7
Demonstrates a simple use of function space in a datatype definition
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
     8
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
     9
Based upon ex/Term.ML
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    10
*)
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    11
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    12
(** ntree **)
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    13
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    14
Goal "ntree(A) = A * (\\<Union>n \\<in> nat. n -> ntree(A))";
529
f0d16216e394 ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
lcp
parents: 515
diff changeset
    15
let open ntree;  val rew = rewrite_rule con_defs in  
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    16
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
    17
end;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    18
qed "ntree_unfold";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    19
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    20
(*A nicer induction rule than the standard one*)
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    21
val major::prems = goal Ntree.thy
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    22
    "[| t \\<in> ntree(A);                                                    \
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    23
\       !!x n h. [| x \\<in> A;  n \\<in> nat;  h \\<in> n -> ntree(A);  \\<forall>i \\<in> n. P(h`i)  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    24
\                |] ==> P(Branch(x,h))                                  \
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    25
\    |] ==> P(t)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    26
by (rtac (major RS ntree.induct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    27
by (etac UN_E 1);
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    28
by (REPEAT_SOME (ares_tac prems));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    29
by (fast_tac (claset() addEs [fun_weaken_type]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    30
by (fast_tac (claset() addDs [apply_type]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    31
qed "ntree_induct";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    32
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    33
(*Induction on ntree(A) to prove an equation*)
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    34
val major::prems = goal Ntree.thy
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    35
  "[| t \\<in> ntree(A);  f \\<in> ntree(A)->B;  g \\<in> ntree(A)->B;                      \
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    36
\     !!x n h. [| x \\<in> A;  n \\<in> nat;  h \\<in> n -> ntree(A);  f O h = g O h |] ==> \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    37
\              f ` Branch(x,h) = g ` Branch(x,h)                          \
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    38
\  |] ==> f`t=g`t";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    39
by (rtac (major RS ntree_induct) 1);
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    40
by (REPEAT_SOME (ares_tac prems));
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    41
by (cut_facts_tac prems 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    42
by (rtac fun_extension 1);
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    43
by (REPEAT_SOME (ares_tac [comp_fun]));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    44
by (asm_simp_tac (simpset() addsimps [comp_fun_apply]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    45
qed "ntree_induct_eqn";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    46
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    47
(**  Lemmas to justify using "Ntree" in other recursive type definitions **)
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    48
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    49
Goalw ntree.defs "A \\<subseteq> B ==> ntree(A) \\<subseteq> ntree(B)";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    50
by (rtac lfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    51
by (REPEAT (rtac ntree.bnd_mono 1));
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    52
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    53
qed "ntree_mono";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    54
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    55
(*Easily provable by induction also*)
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    56
Goalw (ntree.defs@ntree.con_defs) "ntree(univ(A)) \\<subseteq> univ(A)";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    57
by (rtac lfp_lowerbound 1);
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    58
by (rtac (A_subset_univ RS univ_mono) 2);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    59
by Safe_tac;
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    60
by (REPEAT (ares_tac [Pair_in_univ, nat_fun_univ RS subsetD] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    61
qed "ntree_univ";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    62
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    63
val ntree_subset_univ = [ntree_mono, ntree_univ] MRS subset_trans |> standard;
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    64
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    65
8125
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    66
(** ntree recursion **)
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    67
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    68
Goal "ntree_rec(b, Branch(x,h)) \
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    69
\     = b(x, h, \\<lambda>i \\<in> domain(h). ntree_rec(b, h`i))";
8125
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    70
by (rtac (ntree_rec_def RS def_Vrecursor RS trans) 1);
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    71
by (Simp_tac 1);
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    72
by (rewrite_goals_tac ntree.con_defs);
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    73
by (asm_simp_tac (simpset() addsimps [rank_pair2 RSN (2, lt_trans), 
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    74
				      rank_apply]) 1);;
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    75
qed "ntree_rec_Branch";
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    76
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    77
Goalw [ntree_copy_def]
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    78
 "ntree_copy (Branch(x, h)) = Branch(x, \\<lambda>i \\<in> domain(h). ntree_copy (h`i))";
8125
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    79
by (rtac ntree_rec_Branch 1);
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    80
qed "ntree_copy_Branch";
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    81
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    82
Addsimps [ntree_copy_Branch];
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    83
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
    84
Goal "z \\<in> ntree(A) ==> ntree_copy(z) = z";
8125
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    85
by (induct_tac "z" 1);
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    86
by (auto_tac (claset(), simpset() addsimps [domain_of_fun, Pi_Collect_iff]));
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    87
qed "ntree_copy_is_ident";
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    88
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    89
df502e820d07 added recursor
paulson
parents: 6112
diff changeset
    90
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    91
(** maptree **)
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    92
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    93
Goal "maptree(A) = A * (maptree(A) -||> maptree(A))";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    94
let open maptree;  val rew = rewrite_rule con_defs in  
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    95
by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    96
end;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
    97
qed "maptree_unfold";
486
6b58082796f6 Misc minor updates
lcp
parents:
diff changeset
    98
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
    99
(*A nicer induction rule than the standard one*)
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   100
val major::prems = goal Ntree.thy
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
   101
    "[| t \\<in> maptree(A);                                                  \
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
   102
\       !!x n h. [| x \\<in> A;  h \\<in> maptree(A) -||> maptree(A);               \
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
   103
\                   \\<forall>y \\<in> field(h). P(y)                               \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   104
\                |] ==> P(Sons(x,h))                                    \
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   105
\    |] ==> P(t)";
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   106
by (rtac (major RS maptree.induct) 1);
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   107
by (REPEAT_SOME (ares_tac prems));
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   108
by (eresolve_tac [Collect_subset RS FiniteFun_mono1 RS subsetD] 1);
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   109
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   110
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   111
by (Fast_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
   112
qed "maptree_induct";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   113
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   114
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   115
(** maptree2 **)
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   116
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   117
Goal "maptree2(A,B) = A * (B -||> maptree2(A,B))";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   118
let open maptree2;  val rew = rewrite_rule con_defs in  
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   119
by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   120
end;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
   121
qed "maptree2_unfold";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   122
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   123
(*A nicer induction rule than the standard one*)
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   124
val major::prems = goal Ntree.thy
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
   125
    "[| t \\<in> maptree2(A,B);                                                 \
b4e71bd751e4 X-symbols for set theory
paulson
parents: 8125
diff changeset
   126
\       !!x n h. [| x \\<in> A;  h \\<in> B -||> maptree2(A,B);  \\<forall>y \\<in> range(h). P(y) \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   127
\                |] ==> P(Sons2(x,h))                                     \
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   128
\    |] ==> P(t)";
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   129
by (rtac (major RS maptree2.induct) 1);
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   130
by (REPEAT_SOME (ares_tac prems));
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   131
by (eresolve_tac [[subset_refl, Collect_subset] MRS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   132
                  FiniteFun_mono RS subsetD] 1);
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   133
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   134
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   135
by (Fast_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 539
diff changeset
   136
qed "maptree2_induct";
539
01906e4644e0 ZF/ex/Ntree: two new examples, maptree and maptree2
lcp
parents: 529
diff changeset
   137