src/HOL/HOL.thy
changeset 15288 9d49290ed885
parent 15197 19e735596e51
child 15354 9234f5765d9c
equal deleted inserted replaced
15287:55b7f7920622 15288:9d49290ed885
   999 
   999 
  1000 (* The setting up of Quasi_Tac serves as a demo.  Since there is no
  1000 (* The setting up of Quasi_Tac serves as a demo.  Since there is no
  1001    class for quasi orders, the tactics Quasi_Tac.trans_tac and
  1001    class for quasi orders, the tactics Quasi_Tac.trans_tac and
  1002    Quasi_Tac.quasi_tac are not of much use. *)
  1002    Quasi_Tac.quasi_tac are not of much use. *)
  1003 
  1003 
       
  1004 fun decomp_gen sort sign (Trueprop $ t) =
       
  1005   let fun of_sort t = Sign.of_sort sign (type_of t, sort)
       
  1006   fun dec (Const ("Not", _) $ t) = (
       
  1007 	  case dec t of
       
  1008 	    None => None
       
  1009 	  | Some (t1, rel, t2) => Some (t1, "~" ^ rel, t2))
       
  1010 	| dec (Const ("op =",  _) $ t1 $ t2) = 
       
  1011 	    if of_sort t1
       
  1012 	    then Some (t1, "=", t2)
       
  1013 	    else None
       
  1014 	| dec (Const ("op <=",  _) $ t1 $ t2) = 
       
  1015 	    if of_sort t1
       
  1016 	    then Some (t1, "<=", t2)
       
  1017 	    else None
       
  1018 	| dec (Const ("op <",  _) $ t1 $ t2) = 
       
  1019 	    if of_sort t1
       
  1020 	    then Some (t1, "<", t2)
       
  1021 	    else None
       
  1022 	| dec _ = None
       
  1023   in dec t end;
       
  1024 
  1004 structure Quasi_Tac = Quasi_Tac_Fun (
  1025 structure Quasi_Tac = Quasi_Tac_Fun (
  1005   struct
  1026   struct
  1006     val le_trans = thm "order_trans";
  1027     val le_trans = thm "order_trans";
  1007     val le_refl = thm "order_refl";
  1028     val le_refl = thm "order_refl";
  1008     val eqD1 = thm "order_eq_refl";
  1029     val eqD1 = thm "order_eq_refl";
  1010     val less_reflE = thm "order_less_irrefl" RS thm "notE";
  1031     val less_reflE = thm "order_less_irrefl" RS thm "notE";
  1011     val less_imp_le = thm "order_less_imp_le";
  1032     val less_imp_le = thm "order_less_imp_le";
  1012     val le_neq_trans = thm "order_le_neq_trans";
  1033     val le_neq_trans = thm "order_le_neq_trans";
  1013     val neq_le_trans = thm "order_neq_le_trans";
  1034     val neq_le_trans = thm "order_neq_le_trans";
  1014     val less_imp_neq = thm "less_imp_neq";
  1035     val less_imp_neq = thm "less_imp_neq";
  1015 
       
  1016     fun decomp_gen sort sign (Trueprop $ t) =
       
  1017       let fun of_sort t = Sign.of_sort sign (type_of t, sort)
       
  1018       fun dec (Const ("Not", _) $ t) = (
       
  1019               case dec t of
       
  1020                 None => None
       
  1021               | Some (t1, rel, t2) => Some (t1, "~" ^ rel, t2))
       
  1022             | dec (Const ("op =",  _) $ t1 $ t2) = 
       
  1023                 if of_sort t1
       
  1024                 then Some (t1, "=", t2)
       
  1025                 else None
       
  1026             | dec (Const ("op <=",  _) $ t1 $ t2) = 
       
  1027                 if of_sort t1
       
  1028                 then Some (t1, "<=", t2)
       
  1029                 else None
       
  1030             | dec (Const ("op <",  _) $ t1 $ t2) = 
       
  1031                 if of_sort t1
       
  1032                 then Some (t1, "<", t2)
       
  1033                 else None
       
  1034             | dec _ = None
       
  1035       in dec t end;
       
  1036 
       
  1037     val decomp_trans = decomp_gen ["HOL.order"];
  1036     val decomp_trans = decomp_gen ["HOL.order"];
  1038     val decomp_quasi = decomp_gen ["HOL.order"];
  1037     val decomp_quasi = decomp_gen ["HOL.order"];
  1039 
  1038 
  1040   end);  (* struct *)
  1039   end);  (* struct *)
  1041 
  1040 
  1057     val le_trans = thm "order_trans";
  1056     val le_trans = thm "order_trans";
  1058     val le_neq_trans = thm "order_le_neq_trans";
  1057     val le_neq_trans = thm "order_le_neq_trans";
  1059     val neq_le_trans = thm "order_neq_le_trans";
  1058     val neq_le_trans = thm "order_neq_le_trans";
  1060     val less_imp_neq = thm "less_imp_neq";
  1059     val less_imp_neq = thm "less_imp_neq";
  1061     val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq";
  1060     val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq";
  1062 
       
  1063     fun decomp_gen sort sign (Trueprop $ t) =
       
  1064       let fun of_sort t = Sign.of_sort sign (type_of t, sort)
       
  1065       fun dec (Const ("Not", _) $ t) = (
       
  1066               case dec t of
       
  1067                 None => None
       
  1068               | Some (t1, rel, t2) => Some (t1, "~" ^ rel, t2))
       
  1069             | dec (Const ("op =",  _) $ t1 $ t2) = 
       
  1070                 if of_sort t1
       
  1071                 then Some (t1, "=", t2)
       
  1072                 else None
       
  1073             | dec (Const ("op <=",  _) $ t1 $ t2) = 
       
  1074                 if of_sort t1
       
  1075                 then Some (t1, "<=", t2)
       
  1076                 else None
       
  1077             | dec (Const ("op <",  _) $ t1 $ t2) = 
       
  1078                 if of_sort t1
       
  1079                 then Some (t1, "<", t2)
       
  1080                 else None
       
  1081             | dec _ = None
       
  1082       in dec t end;
       
  1083 
       
  1084     val decomp_part = decomp_gen ["HOL.order"];
  1061     val decomp_part = decomp_gen ["HOL.order"];
  1085     val decomp_lin = decomp_gen ["HOL.linorder"];
  1062     val decomp_lin = decomp_gen ["HOL.linorder"];
  1086 
  1063 
  1087   end);  (* struct *)
  1064   end);  (* struct *)
  1088 
  1065