149 val [x, c] = [("x", ty), ("c", ty --> @{typ bool})] |
149 val [x, c] = [("x", ty), ("c", ty --> @{typ bool})] |
150 |> Variable.variant_frees lthy [rel] |
150 |> Variable.variant_frees lthy [rel] |
151 |> map Free |
151 |> map Free |
152 in |
152 in |
153 lambda c |
153 lambda c |
154 (HOLogic.exists_const ty $ |
154 (HOLogic.exists_const ty $ |
155 lambda x (HOLogic.mk_eq (c, (rel $ x)))) |
155 lambda x (HOLogic.mk_eq (c, (rel $ x)))) |
156 end |
156 end |
157 |
157 |
158 |
158 |
159 (* makes the new type definitions and proves non-emptyness*) |
159 (* makes the new type definitions and proves non-emptyness*) |
160 fun typedef_make (qty_name, rel, ty) lthy = |
160 fun typedef_make (qty_name, rel, ty) lthy = |
161 let |
161 let |
162 val typedef_tac = |
162 val typedef_tac = |
163 EVERY1 [rewrite_goal_tac @{thms mem_def}, |
163 EVERY1 [rewrite_goal_tac @{thms mem_def}, |
164 rtac @{thm exI}, |
164 rtac @{thm exI}, |
165 rtac @{thm exI}, |
165 rtac @{thm exI}, |
166 rtac @{thm refl}] |
166 rtac @{thm refl}] |
167 in |
167 in |
168 LocalTheory.theory_result |
168 LocalTheory.theory_result |
169 (Typedef.add_typedef false NONE |
169 (Typedef.add_typedef false NONE |
170 (qty_name, map fst (Term.add_tfreesT ty []), NoSyn) |
170 (qty_name, map fst (Term.add_tfreesT ty []), NoSyn) |
821 apply (simp add: QUOT_TYPE_I_fset.thm10) |
821 apply (simp add: QUOT_TYPE_I_fset.thm10) |
822 done |
822 done |
823 |
823 |
824 ML {* |
824 ML {* |
825 fun mk_rep_abs x = @{term REP_fset} $ (@{term ABS_fset} $ x) |
825 fun mk_rep_abs x = @{term REP_fset} $ (@{term ABS_fset} $ x) |
826 val consts = [@{const_name "Nil"}, @{const_name "append"}, |
826 val consts = [@{const_name "Nil"}, @{const_name "append"}, |
827 @{const_name "Cons"}, @{const_name "membship"}, |
827 @{const_name "Cons"}, @{const_name "membship"}, |
828 @{const_name "card1"}] |
828 @{const_name "card1"}] |
829 *} |
829 *} |
830 |
830 |
831 ML {* |
831 ML {* |
832 fun build_goal ctxt thm constructors qty mk_rep_abs = |
832 fun build_goal ctxt thm constructors qty mk_rep_abs = |
833 let |
833 let |
834 fun is_const (Const (x, t)) = x mem constructors |
834 fun is_constructor (Const (x, _)) = member (op =) constructors x |
|
835 | is_constructor _ = false; |
|
836 |
|
837 fun maybe_mk_rep_abs t = |
|
838 let |
|
839 val _ = writeln ("Maybe: " ^ Syntax.string_of_term ctxt t) |
|
840 in |
|
841 if type_of t = qty then mk_rep_abs t else t |
|
842 end; |
|
843 |
|
844 fun build_aux ctxt1 tm = |
|
845 let |
|
846 val (head, args) = Term.strip_comb tm; |
|
847 val args' = map (build_aux ctxt1) args; |
|
848 in |
|
849 (case head of |
|
850 Abs (x, T, t) => |
|
851 let |
|
852 val ([x'], ctxt2) = Variable.variant_fixes [x] ctxt1; |
|
853 val v = Free (x', T); |
|
854 val t' = subst_bound (v, t); |
|
855 val rec_term = build_aux ctxt2 t'; |
|
856 in Term.lambda_name (x, v) rec_term end |
|
857 | _ => |
|
858 if is_constructor head then |
|
859 maybe_mk_rep_abs (list_comb (head, map maybe_mk_rep_abs args')) |
|
860 else list_comb (head, args')) |
|
861 end; |
|
862 |
|
863 val concl2 = Thm.prop_of thm; |
|
864 in |
|
865 Logic.mk_equals (build_aux ctxt concl2, concl2) |
|
866 end |
|
867 *} |
|
868 |
|
869 ML {* |
|
870 fun build_goal' ctxt thm constructors qty mk_rep_abs = |
|
871 let |
|
872 fun is_const (Const (x, t)) = member (op =) constructors x |
835 | is_const _ = false |
873 | is_const _ = false |
836 |
874 |
837 fun maybe_mk_rep_abs t = |
875 fun maybe_mk_rep_abs t = |
838 let |
876 let |
839 val _ = writeln ("Maybe: " ^ Syntax.string_of_term ctxt t) |
877 val _ = writeln ("Maybe: " ^ Syntax.string_of_term ctxt t) |
840 in |
878 in |
841 if type_of t = qty then mk_rep_abs t else t |
879 if type_of t = qty then mk_rep_abs t else t |
842 end |
880 end |
843 handle TYPE _ => t |
881 (* handle TYPE _ => t*) |
844 |
882 fun build_aux ctxt1 (Abs (x, T, t)) = |
845 fun build_aux (Abs (s, t, tr)) = (Abs (s, t, build_aux tr)) |
883 let |
846 | build_aux (f $ a) = |
884 val ([x'], ctxt2) = Variable.variant_fixes [x] ctxt1; |
|
885 val v = Free (x', T); |
|
886 val t' = subst_bound (v, t); |
|
887 val rec_term = build_aux ctxt2 t'; |
|
888 in Term.lambda_name (x, v) rec_term end |
|
889 | build_aux ctxt1 (f $ a) = |
847 let |
890 let |
848 val (f, args) = strip_comb (f $ a) |
891 val (f, args) = strip_comb (f $ a) |
849 val _ = writeln (Syntax.string_of_term ctxt f) |
892 val _ = writeln (Syntax.string_of_term ctxt f) |
850 in |
893 in |
851 (if is_const f then maybe_mk_rep_abs (list_comb (f, (map maybe_mk_rep_abs (map build_aux args)))) |
894 if is_const f then |
852 else list_comb ((build_aux f), (map build_aux args))) |
895 maybe_mk_rep_abs |
|
896 (list_comb (f, map maybe_mk_rep_abs (map (build_aux ctxt1) args))) |
|
897 else list_comb (build_aux ctxt1 f, map (build_aux ctxt1) args) |
853 end |
898 end |
854 | build_aux x = |
899 | build_aux _ x = |
855 if is_const x then maybe_mk_rep_abs x else x |
900 if is_const x then maybe_mk_rep_abs x else x |
856 |
901 |
857 val concl2 = prop_of thm |
902 val concl2 = term_of (#prop (crep_thm thm)) |
858 in |
903 in |
859 Logic.mk_equals ((build_aux concl2), concl2) |
904 Logic.mk_equals (build_aux ctxt concl2, concl2) |
860 end *} |
905 end *} |
861 |
906 |
862 thm EMPTY_def IN_def UNION_def |
|
863 |
|
864 ML {* val emptyt = symmetric (unlam_def @{context} @{thm EMPTY_def}) *} |
|
865 ML {* val in_t = symmetric (unlam_def @{context} @{thm IN_def}) *} |
|
866 ML {* val uniont = symmetric (unlam_def @{context} @{thm UNION_def}) *} |
|
867 ML {* val cardt = symmetric (unlam_def @{context} @{thm card_def}) *} |
|
868 ML {* val insertt = symmetric (unlam_def @{context} @{thm INSERT_def}) *} |
|
869 ML {* val fset_defs = @{thms EMPTY_def IN_def UNION_def card_def INSERT_def} *} |
907 ML {* val fset_defs = @{thms EMPTY_def IN_def UNION_def card_def INSERT_def} *} |
870 ML {* val fset_defs_sym = [emptyt, in_t, uniont, cardt, insertt] *} |
908 ML {* val fset_defs_sym = map (fn t => symmetric (unlam_def @{context} t)) fset_defs *} |
871 |
909 |
872 ML {* |
910 ML {* |
873 fun dest_cbinop t = |
911 fun dest_cbinop t = |
874 let |
912 let |
875 val (t2, rhs) = Thm.dest_comb t; |
913 val (t2, rhs) = Thm.dest_comb t; |
898 Drule.instantiate' [SOME cT,SOME cT,SOME @{ctyp bool}] [NONE,NONE,NONE,NONE,SOME (@{cpat "op ="})] @{thm arg_cong2} |
936 Drule.instantiate' [SOME cT,SOME cT,SOME @{ctyp bool}] [NONE,NONE,NONE,NONE,SOME (@{cpat "op ="})] @{thm arg_cong2} |
899 ) |
937 ) |
900 *} |
938 *} |
901 |
939 |
902 ML {* |
940 ML {* |
903 fun foo_conv t = |
941 fun split_binop_conv t = |
904 let |
942 let |
905 val (lhs, rhs) = dest_ceq t; |
943 val (lhs, rhs) = dest_ceq t; |
906 val (bop, _) = dest_cbinop lhs; |
944 val (bop, _) = dest_cbinop lhs; |
907 val [clT, cr2] = bop |> Thm.ctyp_of_term |> Thm.dest_ctyp; |
945 val [clT, cr2] = bop |> Thm.ctyp_of_term |> Thm.dest_ctyp; |
908 val [cmT, crT] = Thm.dest_ctyp cr2; |
946 val [cmT, crT] = Thm.dest_ctyp cr2; |
909 in |
947 in |
910 Drule.instantiate' [SOME clT,SOME cmT,SOME crT] [NONE,NONE,NONE,NONE,SOME bop] @{thm arg_cong2} |
948 Drule.instantiate' [SOME clT, SOME cmT, SOME crT] [NONE, NONE, NONE, NONE, SOME bop] @{thm arg_cong2} |
911 end |
949 end |
912 *} |
950 *} |
913 |
951 |
914 ML {* |
952 ML {* |
915 fun foo_tac n thm = |
953 fun split_arg_conv t = |
|
954 let |
|
955 val (lhs, rhs) = dest_ceq t; |
|
956 val (lop, larg) = Thm.dest_comb lhs; |
|
957 val [caT, crT] = lop |> Thm.ctyp_of_term |> Thm.dest_ctyp; |
|
958 in |
|
959 Drule.instantiate' [SOME caT, SOME crT] [NONE, NONE, SOME lop] @{thm arg_cong} |
|
960 end |
|
961 *} |
|
962 |
|
963 ML {* |
|
964 fun split_binop_tac n thm = |
916 let |
965 let |
917 val concl = Thm.cprem_of thm n; |
966 val concl = Thm.cprem_of thm n; |
918 val (_, cconcl) = Thm.dest_comb concl; |
967 val (_, cconcl) = Thm.dest_comb concl; |
919 val rewr = foo_conv cconcl; |
968 val rewr = split_binop_conv cconcl; |
920 (* val _ = tracing (Display.string_of_thm @{context} rewr) |
|
921 val _ = tracing (Display.string_of_thm @{context} thm)*) |
|
922 in |
969 in |
923 rtac rewr n thm |
970 rtac rewr n thm |
924 end |
971 end |
925 handle CTERM _ => Seq.empty |
972 handle CTERM _ => Seq.empty |
926 *} |
973 *} |
927 |
974 |
|
975 ML {* |
|
976 fun split_arg_tac n thm = |
|
977 let |
|
978 val concl = Thm.cprem_of thm n; |
|
979 val (_, cconcl) = Thm.dest_comb concl; |
|
980 val rewr = split_arg_conv cconcl; |
|
981 in |
|
982 rtac rewr n thm |
|
983 end |
|
984 handle CTERM _ => Seq.empty |
|
985 *} |
|
986 |
928 (* Has all the theorems about fset plugged in. These should be parameters to the tactic *) |
987 (* Has all the theorems about fset plugged in. These should be parameters to the tactic *) |
929 |
988 |
930 lemma trueprop_cong: |
989 lemma trueprop_cong: |
931 shows "(a \<equiv> b) \<Longrightarrow> (Trueprop a \<equiv> Trueprop b)" |
990 shows "(a \<equiv> b) \<Longrightarrow> (Trueprop a \<equiv> Trueprop b)" |
932 by auto |
991 by auto |
933 |
992 |
|
993 |
|
994 thm QUOT_TYPE_I_fset.R_trans2 |
934 ML {* |
995 ML {* |
935 fun foo_tac' ctxt = |
996 fun foo_tac' ctxt = |
936 REPEAT_ALL_NEW (FIRST' [ |
997 REPEAT_ALL_NEW (FIRST' [ |
937 rtac @{thm trueprop_cong}, |
998 (* rtac @{thm trueprop_cong},*) |
938 rtac @{thm list_eq_refl}, |
999 rtac @{thm list_eq_refl}, |
939 rtac @{thm cons_preserves}, |
1000 rtac @{thm cons_preserves}, |
940 rtac @{thm mem_respects}, |
1001 rtac @{thm mem_respects}, |
|
1002 rtac @{thm card1_rsp}, |
941 rtac @{thm QUOT_TYPE_I_fset.R_trans2}, |
1003 rtac @{thm QUOT_TYPE_I_fset.R_trans2}, |
942 CHANGED o (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms QUOT_TYPE_I_fset.REP_ABS_rsp})), |
1004 CHANGED o (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms QUOT_TYPE_I_fset.REP_ABS_rsp})), |
943 foo_tac, |
1005 DatatypeAux.cong_tac, |
|
1006 rtac @{thm ext}, |
|
1007 rtac @{thm eq_reflection}, |
944 CHANGED o (ObjectLogic.full_atomize_tac) |
1008 CHANGED o (ObjectLogic.full_atomize_tac) |
945 ]) |
1009 ]) |
946 *} |
1010 *} |
947 |
1011 |
948 ML {* |
1012 ML {* |
949 val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm m1})) |
1013 val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m1})) |
950 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1014 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
951 val cgoal = cterm_of @{theory} goal |
1015 val cgoal = cterm_of @{theory} goal |
952 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
1016 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
953 *} |
1017 *} |
954 |
1018 |
955 (*notation ( output) "prop" ("#_" [1000] 1000) *) |
1019 (*notation ( output) "prop" ("#_" [1000] 1000) *) |
956 notation ( output) "Trueprop" ("#_" [1000] 1000) |
1020 notation ( output) "Trueprop" ("#_" [1000] 1000) |
|
1021 |
|
1022 lemma atomize_eqv [atomize]: "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)" |
|
1023 (is "?rhs \<equiv> ?lhs") |
|
1024 proof |
|
1025 assume "PROP ?lhs" then show "PROP ?rhs" by unfold |
|
1026 next |
|
1027 assume *: "PROP ?rhs" |
|
1028 have "A = B" |
|
1029 proof (cases A) |
|
1030 case True |
|
1031 with * have B by unfold |
|
1032 with `A` show ?thesis by simp |
|
1033 next |
|
1034 case False |
|
1035 with * have "~ B" by auto |
|
1036 with `~ A` show ?thesis by simp |
|
1037 qed |
|
1038 then show "PROP ?lhs" by (rule eq_reflection) |
|
1039 qed |
|
1040 |
957 |
1041 |
958 prove {* (Thm.term_of cgoal2) *} |
1042 prove {* (Thm.term_of cgoal2) *} |
959 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
1043 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
960 apply (tactic {* foo_tac' @{context} 1 *}) |
1044 apply (tactic {* foo_tac' @{context} 1 *}) |
961 done |
1045 done |
962 |
1046 |
963 thm length_append (* Not true but worth checking that the goal is correct *) |
1047 thm length_append (* Not true but worth checking that the goal is correct *) |
964 ML {* |
1048 ML {* |
965 val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm length_append})) |
1049 val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm length_append})) |
966 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1050 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
967 val cgoal = cterm_of @{theory} goal |
1051 val cgoal = cterm_of @{theory} goal |
968 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
1052 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
969 *} |
1053 *} |
970 prove {* Thm.term_of cgoal2 *} |
1054 prove {* Thm.term_of cgoal2 *} |
1057 build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1140 build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1058 ) |
1141 ) |
1059 *} |
1142 *} |
1060 ML {* |
1143 ML {* |
1061 val cgoal = cterm_of @{theory} goal |
1144 val cgoal = cterm_of @{theory} goal |
1062 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
1145 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal) |
1063 *} |
1146 *} |
1064 ML fset_defs_sym |
1147 ML fset_defs_sym |
1065 prove {* (Thm.term_of cgoal2) *} |
1148 prove {* (Thm.term_of cgoal2) *} |
1066 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
1149 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
1067 apply (atomize(full)) |
1150 apply (tactic {* foo_tac' @{context} 1 *}) |
1068 apply (rule_tac trueprop_cong) |
|
1069 apply (atomize(full)) |
|
1070 apply (tactic {* foo_tac' @{context} 1 *}) |
|
1071 apply (rule_tac f = "P" in arg_cong) |
|
1072 sorry |
1151 sorry |
1073 |
1152 |
|
1153 ML {* |
|
1154 fun lift_theorem_fset_aux thm lthy = |
|
1155 let |
|
1156 val ((_, [novars]), lthy2) = Variable.import true [thm] lthy; |
|
1157 val goal = build_goal @{context} novars consts @{typ "'a list"} mk_rep_abs; |
|
1158 val cgoal = cterm_of @{theory} goal; |
|
1159 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal); |
|
1160 val tac = (LocalDefs.unfold_tac @{context} fset_defs) THEN (foo_tac' @{context}) 1; |
|
1161 val cthm = Goal.prove_internal [] cgoal2 (fn _ => tac); |
|
1162 val nthm = MetaSimplifier.rewrite_rule [symmetric cthm] (snd (no_vars (Context.Theory @{theory}, thm))) |
|
1163 val nthm2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same QUOT_TYPE_I_fset.thm10} nthm; |
|
1164 val [nthm3] = ProofContext.export lthy2 lthy [nthm2] |
|
1165 in |
|
1166 nthm3 |
|
1167 end |
|
1168 *} |
|
1169 |
|
1170 ML {* lift_theorem_fset_aux @{thm m1} @{context} *} |
|
1171 |
|
1172 ML {* |
|
1173 fun lift_theorem_fset name thm lthy = |
|
1174 let |
|
1175 val lifted_thm = lift_theorem_fset_aux thm lthy; |
|
1176 val (_, lthy2) = note_thm (name, lifted_thm) lthy; |
|
1177 in |
|
1178 lthy2 |
|
1179 end; |
|
1180 *} |
|
1181 |
|
1182 local_setup {* lift_theorem_fset @{binding "m1_lift"} @{thm m1} *} |
|
1183 local_setup {* lift_theorem_fset @{binding "leqi4_lift"} @{thm list_eq.intros(4)} *} |
|
1184 local_setup {* lift_theorem_fset @{binding "leqi5_lift"} @{thm list_eq.intros(5)} *} |
|
1185 local_setup {* lift_theorem_fset @{binding "m2_lift"} @{thm m2} *} |
|
1186 |
|
1187 thm m1_lift |
|
1188 thm leqi4_lift |
|
1189 thm leqi5_lift |
|
1190 thm m2_lift |
|
1191 |
|
1192 ML Drule.instantiate' |
|
1193 text {* |
|
1194 We lost the schematic variable again :(. |
|
1195 Another variable export will still be necessary... |
|
1196 *} |
|
1197 ML {* |
|
1198 Toplevel.program (fn () => |
|
1199 MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.thm10} ( |
|
1200 Drule.instantiate' [] [NONE, NONE, SOME (@{cpat "REP_fset x"})] @{thm m2_lift} |
|
1201 ) |
|
1202 ) |
|
1203 *} |
|
1204 |
|
1205 thm leqi4_lift |
|
1206 ML {* |
|
1207 val (nam, typ) = (hd (Term.add_vars (term_of (#prop (crep_thm @{thm leqi4_lift}))) [])) |
|
1208 val (_, l) = dest_Type typ |
|
1209 val t = Type ("QuotMain.fset", l) |
|
1210 val v = Var (nam, t) |
|
1211 val cv = cterm_of @{theory} ((term_of @{cpat "REP_fset"}) $ v) |
|
1212 *} |
|
1213 |
|
1214 ML {* |
|
1215 term_of (#prop (crep_thm @{thm sym})) |
|
1216 *} |
|
1217 |
|
1218 ML {* |
|
1219 Toplevel.program (fn () => |
|
1220 MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.thm10} ( |
|
1221 Drule.instantiate' [] [NONE, SOME (cv)] @{thm leqi4_lift} |
|
1222 ) |
|
1223 ) |
|
1224 *} |
|
1225 |
|
1226 |
|
1227 |
|
1228 |
|
1229 |
|
1230 ML {* MRS *} |
1074 thm card1_suc |
1231 thm card1_suc |
1075 |
1232 |
1076 ML {* |
1233 ML {* |
1077 val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm card1_suc})) |
1234 val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm card1_suc})) |
1078 *} |
1235 *} |
1079 ML {* |
1236 ML {* |
1080 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1237 val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs |
1081 *} |
1238 *} |
1082 ML {* |
1239 ML {* |
1083 val cgoal = cterm_of @{theory} goal |
1240 val cgoal = cterm_of @{theory} goal |
1084 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal) |
1241 val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal) |
1085 *} |
1242 *} |
|
1243 ML {* @{term "\<exists>x. P x"} *} |
|
1244 ML {* Thm.bicompose *} |
1086 prove {* (Thm.term_of cgoal2) *} |
1245 prove {* (Thm.term_of cgoal2) *} |
1087 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
1246 apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} ) |
1088 |
1247 apply (tactic {* foo_tac' @{context} 1 *}) |
1089 |
1248 |
1090 |
1249 |
1091 (* |
1250 (* |
1092 datatype obj1 = |
1251 datatype obj1 = |
1093 OVAR1 "string" |
1252 OVAR1 "string" |