diff -r a80d8ec6c998 -r 3dda49e08b9d src/Sequents/ILL.thy --- a/src/Sequents/ILL.thy Fri Jan 04 21:49:06 2019 +0100 +++ b/src/Sequents/ILL.thy Fri Jan 04 23:22:53 2019 +0100 @@ -34,15 +34,15 @@ "_PromAux" :: "three_seqe" ("promaux {_||_||_}") parse_translation \ - [(@{syntax_const "_Trueprop"}, K (single_tr @{const_syntax Trueprop})), - (@{syntax_const "_Context"}, K (two_seq_tr @{const_syntax Context})), - (@{syntax_const "_PromAux"}, K (three_seq_tr @{const_syntax PromAux}))] + [(\<^syntax_const>\_Trueprop\, K (single_tr \<^const_syntax>\Trueprop\)), + (\<^syntax_const>\_Context\, K (two_seq_tr \<^const_syntax>\Context\)), + (\<^syntax_const>\_PromAux\, K (three_seq_tr \<^const_syntax>\PromAux\))] \ print_translation \ - [(@{const_syntax Trueprop}, K (single_tr' @{syntax_const "_Trueprop"})), - (@{const_syntax Context}, K (two_seq_tr' @{syntax_const "_Context"})), - (@{const_syntax PromAux}, K (three_seq_tr' @{syntax_const "_PromAux"}))] + [(\<^const_syntax>\Trueprop\, K (single_tr' \<^syntax_const>\_Trueprop\)), + (\<^const_syntax>\Context\, K (two_seq_tr' \<^syntax_const>\_Context\)), + (\<^const_syntax>\PromAux\, K (three_seq_tr' \<^syntax_const>\_PromAux\))] \ definition liff :: "[o, o] \ o" (infixr "o-o" 45) @@ -271,14 +271,14 @@ ML \ val safe_pack = - @{context} + \<^context> |> fold_rev Cla.add_safe @{thms conj_lemma ll_mp contrad1 contrad2 mp_rule1 mp_rule2 o_a_rule a_not_a_rule} |> Cla.add_unsafe @{thm aux_impl} |> Cla.get_pack; val power_pack = - Cla.put_pack safe_pack @{context} + Cla.put_pack safe_pack \<^context> |> Cla.add_unsafe @{thm impr_contr_der} |> Cla.get_pack; \