src/ZF/ex/Term.ML
changeset 11316 b4e71bd751e4
parent 8126 6244be18fa55
equal deleted inserted replaced
11315:fbca0f74bcef 11316:b4e71bd751e4
    15 end;
    15 end;
    16 qed "term_unfold";
    16 qed "term_unfold";
    17 
    17 
    18 (*Induction on term(A) followed by induction on List *)
    18 (*Induction on term(A) followed by induction on List *)
    19 val major::prems = Goal
    19 val major::prems = Goal
    20     "[| t: term(A);  \
    20     "[| t \\<in> term(A);  \
    21 \       !!x.      [| x: A |] ==> P(Apply(x,Nil));  \
    21 \       !!x.      [| x \\<in> A |] ==> P(Apply(x,Nil));  \
    22 \       !!x z zs. [| x: A;  z: term(A);  zs: list(term(A));  P(Apply(x,zs))  \
    22 \       !!x z zs. [| x \\<in> A;  z \\<in> term(A);  zs: list(term(A));  P(Apply(x,zs))  \
    23 \                 |] ==> P(Apply(x, Cons(z,zs)))  \
    23 \                 |] ==> P(Apply(x, Cons(z,zs)))  \
    24 \    |] ==> P(t)";
    24 \    |] ==> P(t)";
    25 by (rtac (major RS term.induct) 1);
    25 by (rtac (major RS term.induct) 1);
    26 by (etac list.induct 1);
    26 by (etac list.induct 1);
    27 by (etac CollectE 2);
    27 by (etac CollectE 2);
    28 by (REPEAT (ares_tac (prems@[list_CollectD]) 1));
    28 by (REPEAT (ares_tac (prems@[list_CollectD]) 1));
    29 qed "term_induct2";
    29 qed "term_induct2";
    30 
    30 
    31 (*Induction on term(A) to prove an equation*)
    31 (*Induction on term(A) to prove an equation*)
    32 val major::prems = Goal
    32 val major::prems = Goal
    33     "[| t: term(A);  \
    33     "[| t \\<in> term(A);  \
    34 \       !!x zs. [| x: A;  zs: list(term(A));  map(f,zs) = map(g,zs) |] ==> \
    34 \       !!x zs. [| x \\<in> A;  zs: list(term(A));  map(f,zs) = map(g,zs) |] ==> \
    35 \               f(Apply(x,zs)) = g(Apply(x,zs))  \
    35 \               f(Apply(x,zs)) = g(Apply(x,zs))  \
    36 \    |] ==> f(t)=g(t)";
    36 \    |] ==> f(t)=g(t)";
    37 by (rtac (major RS term.induct) 1);
    37 by (rtac (major RS term.induct) 1);
    38 by (resolve_tac prems 1);
    38 by (resolve_tac prems 1);
    39 by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1));
    39 by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1));
    40 qed "term_induct_eqn";
    40 qed "term_induct_eqn";
    41 
    41 
    42 (**  Lemmas to justify using "term" in other recursive type definitions **)
    42 (**  Lemmas to justify using "term" in other recursive type definitions **)
    43 
    43 
    44 Goalw term.defs "A<=B ==> term(A) <= term(B)";
    44 Goalw term.defs "A \\<subseteq> B ==> term(A) \\<subseteq> term(B)";
    45 by (rtac lfp_mono 1);
    45 by (rtac lfp_mono 1);
    46 by (REPEAT (rtac term.bnd_mono 1));
    46 by (REPEAT (rtac term.bnd_mono 1));
    47 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
    47 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
    48 qed "term_mono";
    48 qed "term_mono";
    49 
    49 
    50 (*Easily provable by induction also*)
    50 (*Easily provable by induction also*)
    51 Goalw (term.defs@term.con_defs) "term(univ(A)) <= univ(A)";
    51 Goalw (term.defs@term.con_defs) "term(univ(A)) \\<subseteq> univ(A)";
    52 by (rtac lfp_lowerbound 1);
    52 by (rtac lfp_lowerbound 1);
    53 by (rtac (A_subset_univ RS univ_mono) 2);
    53 by (rtac (A_subset_univ RS univ_mono) 2);
    54 by Safe_tac;
    54 by Safe_tac;
    55 by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1));
    55 by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1));
    56 qed "term_univ";
    56 qed "term_univ";
    57 
    57 
    58 val term_subset_univ = 
    58 val term_subset_univ = 
    59     term_mono RS (term_univ RSN (2,subset_trans)) |> standard;
    59     term_mono RS (term_univ RSN (2,subset_trans)) |> standard;
    60 
    60 
    61 Goal "[| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
    61 Goal "[| t \\<in> term(A);  A \\<subseteq> univ(B) |] ==> t \\<in> univ(B)";
    62 by (REPEAT (ares_tac [term_subset_univ RS subsetD] 1));
    62 by (REPEAT (ares_tac [term_subset_univ RS subsetD] 1));
    63 qed "term_into_univ";
    63 qed "term_into_univ";
    64 
    64 
    65 
    65 
    66 (*** term_rec -- by Vset recursion ***)
    66 (*** term_rec -- by Vset recursion ***)
    67 
    67 
    68 (*Lemma: map works correctly on the underlying list of terms*)
    68 (*Lemma: map works correctly on the underlying list of terms*)
    69 Goal "[| l: list(A);  Ord(i) |] ==>  \
    69 Goal "[| l \\<in> list(A);  Ord(i) |] ==>  \
    70 \     rank(l)<i --> map(%z. (lam x:Vset(i).h(x)) ` z, l) = map(h,l)";
    70 \     rank(l)<i --> map(%z. (\\<lambda>x \\<in> Vset(i).h(x)) ` z, l) = map(h,l)";
    71 by (etac list.induct 1);
    71 by (etac list.induct 1);
    72 by (Simp_tac 1);
    72 by (Simp_tac 1);
    73 by (rtac impI 1);
    73 by (rtac impI 1);
    74 by (subgoal_tac "rank(a)<i & rank(l) < i" 1);
    74 by (subgoal_tac "rank(a)<i & rank(l) < i" 1);
    75 by (asm_simp_tac (simpset() addsimps [rank_of_Ord]) 1);
    75 by (asm_simp_tac (simpset() addsimps [rank_of_Ord]) 1);
    76 by (full_simp_tac (simpset() addsimps list.con_defs) 1);
    76 by (full_simp_tac (simpset() addsimps list.con_defs) 1);
    77 by (blast_tac (claset() addDs (rank_rls RL [lt_trans])) 1);
    77 by (blast_tac (claset() addDs (rank_rls RL [lt_trans])) 1);
    78 qed "map_lemma";
    78 qed "map_lemma";
    79 
    79 
    80 (*Typing premise is necessary to invoke map_lemma*)
    80 (*Typing premise is necessary to invoke map_lemma*)
    81 Goal "ts: list(A) ==> \
    81 Goal "ts \\<in> list(A) ==> \
    82 \     term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))";
    82 \     term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))";
    83 by (rtac (term_rec_def RS def_Vrec RS trans) 1);
    83 by (rtac (term_rec_def RS def_Vrec RS trans) 1);
    84 by (rewrite_goals_tac term.con_defs);
    84 by (rewrite_goals_tac term.con_defs);
    85 by (asm_simp_tac (simpset() addsimps [rank_pair2, map_lemma]) 1);;
    85 by (asm_simp_tac (simpset() addsimps [rank_pair2, map_lemma]) 1);;
    86 qed "term_rec";
    86 qed "term_rec";
    87 
    87 
    88 (*Slightly odd typing condition on r in the second premise!*)
    88 (*Slightly odd typing condition on r in the second premise!*)
    89 val major::prems = Goal
    89 val major::prems = Goal
    90     "[| t: term(A);                                     \
    90     "[| t \\<in> term(A);                                     \
    91 \       !!x zs r. [| x: A;  zs: list(term(A));          \
    91 \       !!x zs r. [| x \\<in> A;  zs: list(term(A));          \
    92 \                    r: list(UN t:term(A). C(t)) |]     \
    92 \                    r \\<in> list(\\<Union>t \\<in> term(A). C(t)) |]     \
    93 \                 ==> d(x, zs, r): C(Apply(x,zs))       \
    93 \                 ==> d(x, zs, r): C(Apply(x,zs))       \
    94 \    |] ==> term_rec(t,d) : C(t)";
    94 \    |] ==> term_rec(t,d) \\<in> C(t)";
    95 by (rtac (major RS term.induct) 1);
    95 by (rtac (major RS term.induct) 1);
    96 by (ftac list_CollectD 1);
    96 by (ftac list_CollectD 1);
    97 by (stac term_rec 1);
    97 by (stac term_rec 1);
    98 by (REPEAT (ares_tac prems 1));
    98 by (REPEAT (ares_tac prems 1));
    99 by (etac list.induct 1);
    99 by (etac list.induct 1);
   107 by (rewtac rew);
   107 by (rewtac rew);
   108 by (rtac (tslist RS term_rec) 1);
   108 by (rtac (tslist RS term_rec) 1);
   109 qed "def_term_rec";
   109 qed "def_term_rec";
   110 
   110 
   111 val prems = Goal
   111 val prems = Goal
   112     "[| t: term(A);                                          \
   112     "[| t \\<in> term(A);                                          \
   113 \       !!x zs r. [| x: A;  zs: list(term(A));  r: list(C) |]  \
   113 \       !!x zs r. [| x \\<in> A;  zs: list(term(A));  r \\<in> list(C) |]  \
   114 \                 ==> d(x, zs, r): C                 \
   114 \                 ==> d(x, zs, r): C                 \
   115 \    |] ==> term_rec(t,d) : C";
   115 \    |] ==> term_rec(t,d) \\<in> C";
   116 by (REPEAT (ares_tac (term_rec_type::prems) 1));
   116 by (REPEAT (ares_tac (term_rec_type::prems) 1));
   117 by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1);
   117 by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1);
   118 qed "term_rec_simple_type";
   118 qed "term_rec_simple_type";
   119 
   119 
   120 AddTCs [term_rec_simple_type];
   120 AddTCs [term_rec_simple_type];
   122 (** term_map **)
   122 (** term_map **)
   123 
   123 
   124 bind_thm ("term_map", term_map_def RS def_term_rec);
   124 bind_thm ("term_map", term_map_def RS def_term_rec);
   125 
   125 
   126 val prems = Goalw [term_map_def]
   126 val prems = Goalw [term_map_def]
   127     "[| t: term(A);  !!x. x: A ==> f(x): B |] ==> term_map(f,t) : term(B)";
   127     "[| t \\<in> term(A);  !!x. x \\<in> A ==> f(x): B |] ==> term_map(f,t) \\<in> term(B)";
   128 by (REPEAT (ares_tac ([term_rec_simple_type, term.Apply_I] @ prems) 1));
   128 by (REPEAT (ares_tac ([term_rec_simple_type, term.Apply_I] @ prems) 1));
   129 qed "term_map_type";
   129 qed "term_map_type";
   130 
   130 
   131 Goal "t: term(A) ==> term_map(f,t) : term({f(u). u:A})";
   131 Goal "t \\<in> term(A) ==> term_map(f,t) \\<in> term({f(u). u \\<in> A})";
   132 by (etac term_map_type 1);
   132 by (etac term_map_type 1);
   133 by (etac RepFunI 1);
   133 by (etac RepFunI 1);
   134 qed "term_map_type2";
   134 qed "term_map_type2";
   135 
   135 
   136 
   136 
   137 (** term_size **)
   137 (** term_size **)
   138 
   138 
   139 bind_thm ("term_size", term_size_def RS def_term_rec);
   139 bind_thm ("term_size", term_size_def RS def_term_rec);
   140 
   140 
   141 Goalw [term_size_def] "t: term(A) ==> term_size(t) : nat";
   141 Goalw [term_size_def] "t \\<in> term(A) ==> term_size(t) \\<in> nat";
   142 by Auto_tac;
   142 by Auto_tac;
   143 qed "term_size_type";
   143 qed "term_size_type";
   144 
   144 
   145 
   145 
   146 (** reflect **)
   146 (** reflect **)
   147 
   147 
   148 bind_thm ("reflect", reflect_def RS def_term_rec);
   148 bind_thm ("reflect", reflect_def RS def_term_rec);
   149 
   149 
   150 Goalw [reflect_def] "t: term(A) ==> reflect(t) : term(A)";
   150 Goalw [reflect_def] "t \\<in> term(A) ==> reflect(t) \\<in> term(A)";
   151 by Auto_tac;
   151 by Auto_tac;
   152 qed "reflect_type";
   152 qed "reflect_type";
   153 
   153 
   154 (** preorder **)
   154 (** preorder **)
   155 
   155 
   156 bind_thm ("preorder", preorder_def RS def_term_rec);
   156 bind_thm ("preorder", preorder_def RS def_term_rec);
   157 
   157 
   158 Goalw [preorder_def] "t: term(A) ==> preorder(t) : list(A)";
   158 Goalw [preorder_def] "t \\<in> term(A) ==> preorder(t) \\<in> list(A)";
   159 by Auto_tac;
   159 by Auto_tac;
   160 qed "preorder_type";
   160 qed "preorder_type";
   161 
   161 
   162 (** postorder **)
   162 (** postorder **)
   163 
   163 
   164 bind_thm ("postorder", postorder_def RS def_term_rec);
   164 bind_thm ("postorder", postorder_def RS def_term_rec);
   165 
   165 
   166 Goalw [postorder_def] "t: term(A) ==> postorder(t) : list(A)";
   166 Goalw [postorder_def] "t \\<in> term(A) ==> postorder(t) \\<in> list(A)";
   167 by Auto_tac;
   167 by Auto_tac;
   168 qed "postorder_type";
   168 qed "postorder_type";
   169 
   169 
   170 
   170 
   171 (** Term simplification **)
   171 (** Term simplification **)
   179 
   179 
   180 (** theorems about term_map **)
   180 (** theorems about term_map **)
   181 
   181 
   182 Addsimps [thm "List.map_compose"];  (*there is also TF.map_compose*)
   182 Addsimps [thm "List.map_compose"];  (*there is also TF.map_compose*)
   183 
   183 
   184 Goal "t: term(A) ==> term_map(%u. u, t) = t";
   184 Goal "t \\<in> term(A) ==> term_map(%u. u, t) = t";
   185 by (etac term_induct_eqn 1);
   185 by (etac term_induct_eqn 1);
   186 by (Asm_simp_tac 1);
   186 by (Asm_simp_tac 1);
   187 qed "term_map_ident";
   187 qed "term_map_ident";
   188 
   188 
   189 Goal "t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u. f(g(u)), t)";
   189 Goal "t \\<in> term(A) ==> term_map(f, term_map(g,t)) = term_map(%u. f(g(u)), t)";
   190 by (etac term_induct_eqn 1);
   190 by (etac term_induct_eqn 1);
   191 by (Asm_simp_tac 1);
   191 by (Asm_simp_tac 1);
   192 qed "term_map_compose";
   192 qed "term_map_compose";
   193 
   193 
   194 Goal "t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))";
   194 Goal "t \\<in> term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))";
   195 by (etac term_induct_eqn 1);
   195 by (etac term_induct_eqn 1);
   196 by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym]) 1);
   196 by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym]) 1);
   197 qed "term_map_reflect";
   197 qed "term_map_reflect";
   198 
   198 
   199 
   199 
   200 (** theorems about term_size **)
   200 (** theorems about term_size **)
   201 
   201 
   202 Goal "t: term(A) ==> term_size(term_map(f,t)) = term_size(t)";
   202 Goal "t \\<in> term(A) ==> term_size(term_map(f,t)) = term_size(t)";
   203 by (etac term_induct_eqn 1);
   203 by (etac term_induct_eqn 1);
   204 by (Asm_simp_tac 1);
   204 by (Asm_simp_tac 1);
   205 qed "term_size_term_map";
   205 qed "term_size_term_map";
   206 
   206 
   207 Goal "t: term(A) ==> term_size(reflect(t)) = term_size(t)";
   207 Goal "t \\<in> term(A) ==> term_size(reflect(t)) = term_size(t)";
   208 by (etac term_induct_eqn 1);
   208 by (etac term_induct_eqn 1);
   209 by (asm_simp_tac(simpset() addsimps [rev_map_distrib RS sym, list_add_rev]) 1);
   209 by (asm_simp_tac(simpset() addsimps [rev_map_distrib RS sym, list_add_rev]) 1);
   210 qed "term_size_reflect";
   210 qed "term_size_reflect";
   211 
   211 
   212 Goal "t: term(A) ==> term_size(t) = length(preorder(t))";
   212 Goal "t \\<in> term(A) ==> term_size(t) = length(preorder(t))";
   213 by (etac term_induct_eqn 1);
   213 by (etac term_induct_eqn 1);
   214 by (asm_simp_tac (simpset() addsimps [length_flat]) 1);
   214 by (asm_simp_tac (simpset() addsimps [length_flat]) 1);
   215 qed "term_size_length";
   215 qed "term_size_length";
   216 
   216 
   217 
   217 
   218 (** theorems about reflect **)
   218 (** theorems about reflect **)
   219 
   219 
   220 Goal "t: term(A) ==> reflect(reflect(t)) = t";
   220 Goal "t \\<in> term(A) ==> reflect(reflect(t)) = t";
   221 by (etac term_induct_eqn 1);
   221 by (etac term_induct_eqn 1);
   222 by (asm_simp_tac (simpset() addsimps [rev_map_distrib]) 1);
   222 by (asm_simp_tac (simpset() addsimps [rev_map_distrib]) 1);
   223 qed "reflect_reflect_ident";
   223 qed "reflect_reflect_ident";
   224 
   224 
   225 
   225 
   226 (** theorems about preorder **)
   226 (** theorems about preorder **)
   227 
   227 
   228 Goal "t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))";
   228 Goal "t \\<in> term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))";
   229 by (etac term_induct_eqn 1);
   229 by (etac term_induct_eqn 1);
   230 by (asm_simp_tac (simpset() addsimps [map_flat]) 1);
   230 by (asm_simp_tac (simpset() addsimps [map_flat]) 1);
   231 qed "preorder_term_map";
   231 qed "preorder_term_map";
   232 
   232 
   233 Goal "t: term(A) ==> preorder(reflect(t)) = rev(postorder(t))";
   233 Goal "t \\<in> term(A) ==> preorder(reflect(t)) = rev(postorder(t))";
   234 by (etac term_induct_eqn 1);
   234 by (etac term_induct_eqn 1);
   235 by (asm_simp_tac(simpset() addsimps [rev_app_distrib, rev_flat, 
   235 by (asm_simp_tac(simpset() addsimps [rev_app_distrib, rev_flat, 
   236 				     rev_map_distrib RS sym]) 1);
   236 				     rev_map_distrib RS sym]) 1);
   237 qed "preorder_reflect_eq_rev_postorder";
   237 qed "preorder_reflect_eq_rev_postorder";
   238 
   238