src/HOL/Lex/RegSet.ML
author wenzelm
Sun, 12 Jul 1998 11:49:17 +0200
changeset 5132 24f992a25adc
parent 5069 3ea049f7979d
child 6162 484adda70b65
permissions -rw-r--r--
isatool expandshort;

(*  Title:      HOL/Lex/RegSet.ML
    ID:         $Id$
    Author:     Tobias Nipkow
    Copyright   1998 TUM
*)

AddIffs [star.NilI];
Addsimps [star.ConsI];
AddIs [star.ConsI];

Goal "(!xs: set xss. xs:S) --> concat xss : star S";
by (induct_tac "xss" 1);
by (ALLGOALS Asm_full_simp_tac);
qed_spec_mp "concat_in_star";

Goal
 "w : star U = (? us. (!u : set(us). u : U) & (w = concat us))";
by (rtac iffI 1);
 be star.induct 1;
  by (res_inst_tac [("x","[]")] exI 1);
  by (Simp_tac 1);
 by (Clarify_tac 1);
 by (res_inst_tac [("x","a#us")] exI 1);
 by (Asm_simp_tac 1);
by (Clarify_tac 1);
by (asm_simp_tac (simpset() addsimps [concat_in_star]) 1);
qed "in_star";